├── SummerSchoolSophia ├── exercise6_todo.v ├── exercise8_todo.v ├── exercise1_todo.v ├── exercise1_solution.v ├── exercise4_todo.v ├── exercise8_solution.v ├── exercise4_solution.v ├── exercise2_todo.v ├── exercise2_solution.v ├── exercise3_todo.v ├── exercise5_todo.v ├── exercise7_todo.v ├── exercise3_solution.v ├── exercise6_solution.v ├── lesson6.v ├── lesson8.v ├── lesson5.v ├── exercise5_solution.v ├── exercise7_solution.v ├── lesson3.v ├── lesson7.v ├── lesson2.v ├── lesson4.v └── lesson1.v ├── Makefile ├── .gitignore ├── .github └── workflows │ └── docker-action.yml ├── _CoqProject ├── AnIntroductionToSmallScaleReflectionInCoq ├── section7.v ├── section5.v ├── section3.v └── section4.v ├── .travis.yml ├── coq-tutorial_material.opam ├── meta.yml ├── README.md ├── Ssreflect15min └── quesako.v └── AnSsreflectTutorial └── tutorial_Gonthier_LeRoux.v /SummerSchoolSophia/exercise6_todo.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | 3 | Implicit Type p q r : bool. 4 | Implicit Type m n a b c : nat. 5 | 6 | (** *** Exercise 1: 7 | - bla bla 8 | *) 9 | 10 | Lemma bool_gimmics1 a : a != a.-1 -> a != 0. 11 | 12 | 13 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | ifeq "$(COQBIN)" "" 2 | COQBIN=$(dir $(shell which coqtop))/ 3 | endif 4 | 5 | %: Makefile.coq 6 | 7 | Makefile.coq: _CoqProject 8 | $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq 9 | 10 | tests: all 11 | @$(MAKE) -C tests -s clean 12 | @$(MAKE) -C tests -s all 13 | 14 | -include Makefile.coq 15 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | *.d 2 | *.vo 3 | *.vio 4 | *.vos 5 | *.vok 6 | *.cm* 7 | *~ 8 | *.glob 9 | *.aux 10 | *.a 11 | *.o 12 | Make*.coq 13 | Make*.coq.bak 14 | Make*.coq.conf 15 | mathcomp/ssreflect/ssreflect.ml4 16 | mathcomp/ssreflect/ssrmatching.ml4 17 | mathcomp/ssreflect/ssrmatching.mli 18 | mathcomp/ssreflect/ssrmatching.v 19 | mathcomp/ssreflect/ssreflect_plugin.mllib 20 | mathcomp/ssreflect/ssreflect_plugin.mlpack 21 | mathcomp/ssreflect.ml4 22 | mathcomp/ssrmatching.ml4 23 | mathcomp/ssrmatching.mli 24 | mathcomp/ssrmatching.v 25 | mathcomp/ssreflect_plugin.mllib 26 | mathcomp/ssreflect_plugin.mlpack 27 | mathcomp/test_suite/test_hierarchy_all.v 28 | mathcomp/test_suite/*.v.out.new 29 | mathcomp-*.tar.gz 30 | *# 31 | htmldoc/mathcomp.*html 32 | htmldoc/index*.*html 33 | htmldoc/depend 34 | htmldoc/depend.js 35 | htmldoc/*.css 36 | -------------------------------------------------------------------------------- /.github/workflows/docker-action.yml: -------------------------------------------------------------------------------- 1 | # This file was generated from `meta.yml`, please do not edit manually. 2 | # Follow the instructions on https://github.com/coq-community/templates to regenerate. 3 | name: Docker CI 4 | 5 | on: 6 | push: 7 | branches: 8 | - master 9 | pull_request: 10 | branches: 11 | - '**' 12 | 13 | jobs: 14 | build: 15 | # the OS must be GNU/Linux to be able to use the docker-coq-action 16 | runs-on: ubuntu-latest 17 | strategy: 18 | matrix: 19 | image: 20 | - 'mathcomp/mathcomp:2.1.0-coq-8.18' 21 | fail-fast: false 22 | steps: 23 | - uses: actions/checkout@v3 24 | - uses: coq-community/docker-coq-action@v1 25 | with: 26 | opam_file: 'coq-tutorial_material.opam' 27 | custom_image: ${{ matrix.image }} 28 | 29 | 30 | # See also: 31 | # https://github.com/coq-community/docker-coq-action#readme 32 | # https://github.com/erikmd/docker-coq-github-action-demo 33 | -------------------------------------------------------------------------------- /_CoqProject: -------------------------------------------------------------------------------- 1 | -R . tutorial 2 | 3 | -arg -w -arg -notation-overridden 4 | -arg -w -arg -undo-batch-mode 5 | -arg -w -arg -ambiguous-paths 6 | 7 | AnSsreflectTutorial/tutorial_Gonthier_LeRoux.v 8 | AnIntroductionToSmallScaleReflectionInCoq/section3.v 9 | AnIntroductionToSmallScaleReflectionInCoq/section4.v 10 | AnIntroductionToSmallScaleReflectionInCoq/section5.v 11 | AnIntroductionToSmallScaleReflectionInCoq/section6.v 12 | AnIntroductionToSmallScaleReflectionInCoq/section7.v 13 | SummerSchoolSophia/lesson1.v 14 | SummerSchoolSophia/lesson2.v 15 | SummerSchoolSophia/lesson3.v 16 | SummerSchoolSophia/lesson4.v 17 | SummerSchoolSophia/lesson5.v 18 | SummerSchoolSophia/lesson6.v 19 | SummerSchoolSophia/lesson7.v 20 | SummerSchoolSophia/lesson8.v 21 | SummerSchoolSophia/exercise1_solution.v 22 | SummerSchoolSophia/exercise2_solution.v 23 | SummerSchoolSophia/exercise3_solution.v 24 | SummerSchoolSophia/exercise4_solution.v 25 | SummerSchoolSophia/exercise5_solution.v 26 | SummerSchoolSophia/exercise6_solution.v 27 | SummerSchoolSophia/exercise7_solution.v 28 | SummerSchoolSophia/exercise8_solution.v 29 | 30 | Ssreflect15min/quesako.v -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise8_todo.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | 3 | 4 | (** *** Exercise 1: 5 | - Let's define the subtype of odd and even natural numbers 6 | - Intrument Coq to recognize odd/even number built out 7 | of product and successor 8 | - Inherit on [odd_nat] the [eqType] structure 9 | *) 10 | 11 | Structure odd_nat := Odd { 12 | oval :> nat; 13 | oprop : odd oval 14 | }. 15 | Lemma oddP (n : odd_nat) : odd n. 16 | Proof. by case: n. Qed. 17 | 18 | Structure even_nat := Even { 19 | eval :> nat; 20 | eprop : ~~ (odd eval) 21 | }. 22 | Lemma evenP (n : even_nat) : ~~ (odd n). 23 | Proof. by case: n. Qed. 24 | 25 | Example test_odd (n : odd_nat) : ~~ (odd 6) && odd (n * 3). 26 | Proof. Fail by rewrite oddP evenP. Abort. 27 | 28 | Canonical even_0 := Even 0 isT. 29 | 30 | Lemma oddS n : ~~ (odd n) -> odd n.+1. 31 | Proof. 32 | Qed. 33 | 34 | Lemma evenS n : (odd n) -> ~~ (odd n.+1). 35 | Proof. 36 | Qed. 37 | 38 | Canonical odd_even (m : even_nat) := 39 | Odd m.+1 (oddS m (eprop m)). 40 | Canonical even_odd (m : odd_nat) := 41 | 42 | Lemma odd_mulP (n m : odd_nat) : odd (n * m). 43 | Proof. 44 | Qed. 45 | Canonical odd_mul (n m : odd_nat) := 46 | 47 | Example test_odd (n : odd_nat) : ~~ (odd 6) && odd (n * 3). 48 | Proof. by rewrite oddP evenP. Qed. 49 | 50 | Fail Check forall n m : odd_nat, n == m. 51 | 52 | Canonical odd_subType := 53 | Definition odd_eqMixin := 54 | Canonical odd_eqType := 55 | 56 | Check forall n m : odd_nat, n == m. 57 | 58 | -------------------------------------------------------------------------------- /AnIntroductionToSmallScaleReflectionInCoq/section7.v: -------------------------------------------------------------------------------- 1 | (******************************************************************************) 2 | (* Solutions of exercises : Appendix *) 3 | (******************************************************************************) 4 | 5 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq path. 6 | From mathcomp Require Import choice fintype tuple finset. 7 | 8 | Set Implicit Arguments. 9 | Unset Strict Implicit. 10 | Import Prenex Implicits. 11 | 12 | (******************************************************************************) 13 | (* Exercise 7.0.3 *) 14 | (******************************************************************************) 15 | 16 | Search _ "_ * _". 17 | Print muln. 18 | Print muln_rec. 19 | 20 | Search _ "_ ==> _". 21 | Print implb. 22 | Search (_ ==> true). 23 | 24 | Search _ (_ && _ = _ && _). 25 | 26 | (******************************************************************************) 27 | (* Exercise 7.0.4 *) 28 | (******************************************************************************) 29 | 30 | 31 | (* According to the name of the operator and to the table of suffixes *) 32 | (*it should be mulnC *) 33 | Search (_ * _ = _ * _). 34 | (* The above command does not work since commutativity for muln is *) 35 | (*stated using the predicate on binary operations *) 36 | Search _ commutative muln. 37 | 38 | Check orbCA. 39 | Print left_commutative. 40 | -------------------------------------------------------------------------------- /.travis.yml: -------------------------------------------------------------------------------- 1 | dist: trusty 2 | sudo: required 3 | language: generic 4 | 5 | branches: 6 | only: 7 | - master 8 | 9 | services: 10 | - docker 11 | 12 | env: 13 | global: 14 | - NJOBS="2" 15 | - CONTRIB_NAME="tutorial_material" 16 | matrix: 17 | - DOCKERIMAGE="mathcomp/mathcomp:1.13.0-coq-8.14" 18 | 19 | install: | 20 | # Prepare the COQ container 21 | docker pull "${DOCKERIMAGE}" 22 | docker run -d -i --init --name=COQ -v ${TRAVIS_BUILD_DIR}:/home/coq/${CONTRIB_NAME} -w /home/coq/${CONTRIB_NAME} "${DOCKERIMAGE}" 23 | docker exec COQ /bin/bash --login -c " 24 | # This bash script is double-quoted to interpolate Travis CI env vars: 25 | echo \"Build triggered by ${TRAVIS_EVENT_TYPE}\" 26 | export PS4='+ \e[33;1m(\$0 @ line \$LINENO) \$\e[0m ' 27 | set -ex # -e = exit on failure; -x = trace for debug 28 | # Using flambda makes sense here as we usually get ~10% faster 29 | # builds in math-comp. 30 | opam switch \$COMPILER_EDGE 31 | opam update -y 32 | opam config list 33 | opam repo list 34 | opam list 35 | sudo chown -R coq:coq /home/coq/${CONTRIB_NAME} 36 | opam pin add -y -n tutorial_material . 37 | opam install -y -vvv --deps-only tutorial_material 38 | " 39 | 40 | script: 41 | - echo -e "${ANSI_YELLOW}Building ${CONTRIB_NAME}...${ANSI_RESET}" && echo -en 'travis_fold:start:script\\r' 42 | - | 43 | docker exec COQ /bin/bash --login -c " 44 | export PS4='+ \e[33;1m(\$0 @ line \$LINENO) \$\e[0m ' 45 | set -ex 46 | opam install -y -vvv tutorial_material 47 | " 48 | - docker stop COQ # optional 49 | - echo -en 'travis_fold:end:script\\r' 50 | -------------------------------------------------------------------------------- /coq-tutorial_material.opam: -------------------------------------------------------------------------------- 1 | # This file was generated from `meta.yml`, please do not edit manually. 2 | # Follow the instructions on https://github.com/coq-community/templates to regenerate. 3 | 4 | opam-version: "2.0" 5 | maintainer: "thery@sophia.inria.fr" 6 | version: "dev" 7 | 8 | homepage: "https://github.com/math-comp/tutorial_material" 9 | dev-repo: "git+https://github.com/math-comp/tutorial_material.git" 10 | bug-reports: "https://github.com/math-comp/tutorial_material/issues" 11 | license: "MIT" 12 | 13 | synopsis: "Scripts associated to tutorials for mathcomp." 14 | description: """ 15 | 16 | ## Tutorial materials 17 | 18 | Scripts associated to tutorials for mathcomp. 19 | 20 | It contains 21 | - [AnIntroductionToSmallScaleReflectionInCoq](https://github.com/math-comp/tutorial_material/blob/master/AnIntroductionToSmallScaleReflectionInCoq) associated to [An Introduction To Small Scale Reflection In Coq](https://hal.inria.fr/inria-00515548v4/document) 22 | - [AnSsreflectTutorial](https://github.com/math-comp/tutorial_material/blob/master/AnSsreflectTutorial) associated to [An Ssreflect Tutorial](https://hal.inria.fr/inria-00407778) 23 | - [SummerSchoolSophia](https://github.com/math-comp/tutorial_material/tree/master/SummerSchoolSophia) associated to a [5-day School](https://team.inria.fr/marelle/en/coq-winter-school-2018/) on Mathematical Components at Sophia-Antipolis""" 24 | 25 | build: [make "-j%{jobs}%"] 26 | install: [make "install"] 27 | depends: [ 28 | "coq" {(>= "8.18")} 29 | "coq-mathcomp-ssreflect" {(>= "2.1.0")} 30 | ] 31 | 32 | tags: [ 33 | "keyword:mathematical component" 34 | "keyword:tutorial" 35 | "logpath:tutorial_material" 36 | ] 37 | authors: [ 38 | "Laurent Théry" 39 | ] 40 | -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise1_todo.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | Set Implicit Arguments. 3 | Unset Strict Implicit. 4 | Unset Printing Implicit Defensive. 5 | 6 | (** *** Exercise : 7 | - Define the option container with constructors None and Some 8 | - Define the "projection" default 9 | *) 10 | Inductive option 11 | Arguments Some {_}. 12 | Arguments None {_}. 13 | 14 | Definition default 15 | Eval lazy in default 3 None. (* 3 *) 16 | Eval lazy in default 3 (Some 4). (* 4 *) 17 | 18 | (** *** Exercise : 19 | Define boolean negation 20 | *) 21 | Definition negb b := 22 | Notation "~~ x" := (negb x). 23 | 24 | Eval lazy in negb true. 25 | Eval lazy in negb false. 26 | 27 | (** *** Exercise : 28 | Use the [iter] function below to define: 29 | - addition over natural numbers. 30 | - multiplication over natural unmbers. 31 | *) 32 | Fixpoint iter (T : Type) n op (x : T) := 33 | if n is p.+1 then op (iter p op x) else x. 34 | Arguments iter {T}. 35 | 36 | Definition addn n m := 37 | 38 | Eval lazy in addn 3 4. 39 | 40 | Definition muln n m := 41 | 42 | Eval lazy in muln 3 4. 43 | 44 | (** *** Exercise : 45 | - Define muln by recursion 46 | *) 47 | Fixpoint muln_rec n m := 48 | 49 | (** *** Exercise : 50 | - Use the the existing map function to define a functions that adds 2 to 51 | all elements of a list of integers. 52 | - Use the result of the previous exercise and the iter function to define 53 | a function that maps the natural number n to the list containing the n first 54 | odd numbers. (start with the empty list and then at each step, add 1 in front 55 | and increase all other elements by 2).*) 56 | 57 | Definition add2list s := 58 | 59 | Definition odds n := 60 | 61 | -------------------------------------------------------------------------------- /meta.yml: -------------------------------------------------------------------------------- 1 | --- 2 | fullname: tutorial_material 3 | shortname: tutorial_material 4 | organization: math-comp 5 | community: false 6 | dune: false 7 | action: true 8 | 9 | synopsis: Scripts associated to tutorials for mathcomp. 10 | 11 | description: |- 12 | 13 | ## Tutorial materials 14 | 15 | Scripts associated to tutorials for mathcomp. 16 | 17 | It contains 18 | - [AnIntroductionToSmallScaleReflectionInCoq](https://github.com/math-comp/tutorial_material/blob/master/AnIntroductionToSmallScaleReflectionInCoq) associated to [An Introduction To Small Scale Reflection In Coq](https://hal.inria.fr/inria-00515548v4/document) 19 | - [AnSsreflectTutorial](https://github.com/math-comp/tutorial_material/blob/master/AnSsreflectTutorial) associated to [An Ssreflect Tutorial](https://hal.inria.fr/inria-00407778) 20 | - [SummerSchoolSophia](https://github.com/math-comp/tutorial_material/tree/master/SummerSchoolSophia) associated to a [5-day School](https://team.inria.fr/marelle/en/coq-winter-school-2018/) on Mathematical Components at Sophia-Antipolis 21 | 22 | 23 | authors: 24 | - name: Laurent Théry 25 | 26 | maintainers: 27 | - name: Laurent Théry 28 | nickname: thery 29 | 30 | opam-file-maintainer: thery@sophia.inria.fr 31 | 32 | license: 33 | fullname: MIT License 34 | identifier: MIT 35 | 36 | supported_coq_versions: 37 | text: '8.18 or later' 38 | opam: '{(>= "8.18")}' 39 | 40 | dependencies: 41 | - opam: 42 | name: coq-mathcomp-ssreflect 43 | version: '{(>= "2.1.0")}' 44 | description: |- 45 | [MathComp ssreflect 2.1 or later](https://math-comp.github.io) 46 | 47 | tested_coq_opam_versions: 48 | - version: '2.1.0-coq-8.18' 49 | repo: 'mathcomp/mathcomp' 50 | 51 | namespace: tutorial_material 52 | 53 | keywords: 54 | - name: mathematical component 55 | - name: tutorial 56 | 57 | 58 | --- -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise1_solution.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | Set Implicit Arguments. 3 | Unset Strict Implicit. 4 | Unset Printing Implicit Defensive. 5 | 6 | (** *** Exercise : 7 | - Define the option container with constructors None and Some 8 | - Define the "projection" default 9 | *) 10 | Inductive option 11 | (*D*)A := Some (a : A) | None. 12 | Arguments Some {_}. 13 | Arguments None {_}. 14 | 15 | Definition default 16 | (*D*)A (a : A) (x : option A) := if x is Some v then v else a. 17 | Eval lazy in default 3 None. (* 3 *) 18 | Eval lazy in default 3 (Some 4). (* 4 *) 19 | 20 | (** *** Exercise : 21 | Define boolean negation 22 | *) 23 | Definition negb b := 24 | (*D*)if b then false else true. 25 | Notation "~~ x" := (negb x). 26 | 27 | Eval lazy in negb true. 28 | Eval lazy in negb false. 29 | 30 | (** *** Exercise : 31 | Use the [iter] function below to define: 32 | - addition over natural numbers. 33 | - multiplication over natural unmbers. 34 | *) 35 | Fixpoint iter (T : Type) n op (x : T) := 36 | if n is p.+1 then op (iter p op x) else x. 37 | Arguments iter {T}. 38 | 39 | Definition addn n m := 40 | (*D*)iter n S m. 41 | 42 | Eval lazy in addn 3 4. 43 | 44 | Definition muln n m := 45 | (*D*)iter n (addn m) 0. 46 | 47 | Eval lazy in muln 3 4. 48 | 49 | (** *** Exercise : 50 | - Define muln by recursion 51 | *) 52 | Fixpoint muln_rec n m := 53 | (*D*)if n is p.+1 then m + (muln_rec p m) else 0. 54 | 55 | (** *** Exercise : 56 | - Use the the existing map function to define a functions that adds 2 to 57 | all elements of a list of integers. 58 | - Use the result of the previous exercise and the iter function to define 59 | a function that maps the natural number n to the list containing the n first 60 | odd numbers. (start with the empty list and then at each step, add 1 in front 61 | and increase all other elements by 2).*) 62 | 63 | Definition add2list s := 64 | (*D*)map (fun x => x.+2) s. 65 | 66 | Definition odds n := 67 | (*D*) iter n (fun s => 1 :: add2list s) [::]. 68 | 69 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | 5 | # tutorial_material 6 | 7 | [![Docker CI][docker-action-shield]][docker-action-link] 8 | 9 | [docker-action-shield]: https://github.com/math-comp/tutorial_material/workflows/Docker%20CI/badge.svg?branch=master 10 | [docker-action-link]: https://github.com/math-comp/tutorial_material/actions?query=workflow:"Docker%20CI" 11 | 12 | 13 | 14 | 15 | 16 | ## Tutorial materials 17 | 18 | Scripts associated to tutorials for mathcomp. 19 | 20 | It contains 21 | - [AnIntroductionToSmallScaleReflectionInCoq](https://github.com/math-comp/tutorial_material/blob/master/AnIntroductionToSmallScaleReflectionInCoq) associated to [An Introduction To Small Scale Reflection In Coq](https://hal.inria.fr/inria-00515548v4/document) 22 | - [AnSsreflectTutorial](https://github.com/math-comp/tutorial_material/blob/master/AnSsreflectTutorial) associated to [An Ssreflect Tutorial](https://hal.inria.fr/inria-00407778) 23 | - [SummerSchoolSophia](https://github.com/math-comp/tutorial_material/tree/master/SummerSchoolSophia) associated to a [5-day School](https://team.inria.fr/marelle/en/coq-winter-school-2018/) on Mathematical Components at Sophia-Antipolis 24 | 25 | ## Meta 26 | 27 | - Author(s): 28 | - Laurent Théry 29 | - License: [MIT License](LICENSE) 30 | - Compatible Coq versions: 8.18 or later 31 | - Additional dependencies: 32 | - [MathComp ssreflect 2.1 or later](https://math-comp.github.io) 33 | - Coq namespace: `tutorial_material` 34 | - Related publication(s): none 35 | 36 | ## Building and installation instructions 37 | 38 | The easiest way to install the latest released version of tutorial_material 39 | is via [OPAM](https://opam.ocaml.org/doc/Install.html): 40 | 41 | ```shell 42 | opam repo add coq-released https://coq.inria.fr/opam/released 43 | opam install coq-tutorial_material 44 | ``` 45 | 46 | To instead build and install manually, do: 47 | 48 | ``` shell 49 | git clone https://github.com/math-comp/tutorial_material.git 50 | cd tutorial_material 51 | make # or make -j 52 | make install 53 | ``` 54 | 55 | 56 | 57 | -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise4_todo.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | 3 | Implicit Type P Q R : Prop. 4 | 5 | (** *** Exercise 0: 6 | - Define not. In type theory negation is defined in terms 7 | of [False]. 8 | *) 9 | 10 | Definition not P := 11 | . 12 | 13 | (** *** Exercise 1: 14 | - Prove the negation of the excluded middle. 15 | *) 16 | Lemma ex0 P : not (P /\ not P). 17 | Proof. 18 | Qed. 19 | 20 | (** *** Exercise 2: 21 | - Declare iff (the constructor being called [iff_intro]). 22 | - Define iff1 of the given type 23 | *) 24 | Inductive iff P Q := 25 | . 26 | 27 | Definition iff1 P Q : iff P Q -> P -> Q := 28 | 29 | (** *** Exercise 3: 30 | - Declare xor: two constructors, both have two arguments 31 | - Prove the following lemmas 32 | *) 33 | 34 | Inductive xor P Q : Prop := 35 | . 36 | 37 | Lemma xorC P Q : iff (xor P Q) (xor Q P). 38 | Proof. 39 | Qed. 40 | 41 | 42 | Lemma xor1 P Q : (xor P Q) -> not Q -> P. 43 | Proof. 44 | Qed. 45 | 46 | Lemma xor2 P Q Z : (xor P Q) -> (xor Q Z) -> iff P Z. 47 | Proof. 48 | Qed. 49 | 50 | (** *** Exercise 4: 51 | - Declare exists2 52 | - Prove a lemma ex1 -> ex2 T 53 | *) 54 | 55 | Inductive ex2 T (P Q : pred T) : Prop := 56 | . 57 | 58 | Lemma ex2T T (P : pred T) : (exists x, P x) -> (ex2 T P P). 59 | Proof. 60 | Qed. 61 | 62 | (** *** Exercise 5: 63 | - Write the induction principle for lists 64 | *) 65 | Definition induction_seq A (P : seq A -> Prop) : 66 | P nil -> (forall a l, P l -> P (a :: l)) -> forall l, P l := 67 | . 68 | 69 | 70 | (** *** Exercise 6: 71 | - remeber [=> /view] to prove the following lemma 72 | - the two relevant views are [prime_gt1] and [dvdn_leq] 73 | - Note: [=> /view] combines well with [->] (lesson 3) 74 | - Hint: the proof can be a one liner [by move=> ....] 75 | - Recall: the notation "_ < _ <= _" hides a conjunction 76 | *) 77 | About prime_gt1. 78 | About dvdn_leq. 79 | 80 | Lemma ex_view p : prime p -> p %| 7 -> 1 < p <= 7. 81 | Proof. 82 | Qed. 83 | 84 | (** *** Exercise 7: 85 | - Define the indexed data type of Cherry tree: 86 | + the index is a bool and must be truee iff the tree is completely 87 | flourished 88 | + leaves can be either Flower or Bud 89 | + the third constructor is called Node and has two sub trees 90 | *) 91 | Inductive cherryt : bool -> Type := 92 | 93 | Check Node _ Flower Flower : cherryt true. 94 | Check Node _ Bud Bud : cherryt false. 95 | Fail Check Node _ Flower Bud. 96 | 97 | 98 | -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise8_solution.v: -------------------------------------------------------------------------------- 1 | From HB Require Import structures. 2 | From mathcomp Require Import all_ssreflect. 3 | 4 | 5 | (** *** Exercise 1: 6 | - Let's define the subtype of odd and even natural numbers 7 | - Intrument Coq to recognize odd/even number built out 8 | of product and successor 9 | - Inherit on [odd_nat] the [eqType] structure 10 | *) 11 | 12 | Structure odd_nat := Odd { 13 | oval :> nat; 14 | oprop : odd oval 15 | }. 16 | 17 | (* Prove the main property of [odd_nat] *) 18 | Lemma oddP (n : odd_nat) : odd n. 19 | Proof. 20 | (*D*)by case: n. 21 | Qed. 22 | 23 | Structure even_nat := Even { 24 | eval :> nat; 25 | eprop : ~~ (odd eval) 26 | }. 27 | 28 | (* Prove the main property of [even_nat] *) 29 | Lemma evenP (n : even_nat) : ~~ (odd n). 30 | Proof. 31 | (*D*) by case: n. 32 | Qed. 33 | 34 | (* The objective is to make it work: knowing that [n] is [odd] 35 | Coq should infer that [n * 3] is also odd, and that [6] is even *) 36 | Example test_odd (n : odd_nat) : ~~ (odd 6) && odd (n * 3). 37 | Proof. Fail by rewrite oddP evenP. Abort. 38 | 39 | (* Let's start by telling Coq that 0 is even *) 40 | Canonical even_0 : even_nat := Even 0 isT. 41 | 42 | Lemma oddS n : ~~ (odd n) -> odd n.+1. 43 | Proof. 44 | (*D*)by rewrite /=; case: (odd n). 45 | Qed. 46 | 47 | Lemma evenS n : (odd n) -> ~~ (odd n.+1). 48 | Proof. 49 | (*D*)by rewrite /=; case: (odd n). 50 | Qed. 51 | 52 | (* Here we teach Coq that if [m] is even, then [m.+1] is odd *) 53 | Canonical odd_even (m : even_nat) : odd_nat := 54 | Odd m.+1 (oddS m (eprop m)). 55 | 56 | (* Implement the dual, teach coq that if [m] is odd then [m.+1] is even *) 57 | Canonical even_odd (m : odd_nat) : even_nat := 58 | (*D*)Even m.+1 (evenS m (oprop m)). 59 | 60 | (* Now let's deal with multiplication *) 61 | Lemma odd_mulP (n m : odd_nat) : odd (n * m). 62 | Proof. 63 | (*D*)by rewrite oddM !oddP. 64 | Qed. 65 | 66 | (* Teach Coq that [*] preserves being odd *) 67 | Canonical oddM (n m : odd_nat) : odd_nat := 68 | (*D*)Odd (n * m) (odd_mulP n m). 69 | 70 | (* Enjoy! *) 71 | Example test_odd (n : odd_nat) : ~~ (odd 6) && odd (n * 3). 72 | Proof. by rewrite oddP evenP. Qed. 73 | 74 | (* We can't use [==] on odd natural numbers because 75 | [odd_nat] is not an [eqType] *) 76 | Fail Check forall n m : odd_nat, n == m. 77 | 78 | (* Use the subtype machinery (that we used for tuples) in order 79 | to teach Coq that [odd_nat] is an [eqType] *) 80 | HB.instance Definition _ := 81 | (*D*) [isSub for oval]. 82 | HB.instance Definition _ := 83 | (*D*)[Equality of odd_nat by <:]. 84 | 85 | (* Enjoy *) 86 | Check forall n m : odd_nat, n == m. 87 | 88 | (* Now do the same for [even_nat] *) 89 | Fail Check forall (n m : even_nat), m == n. 90 | 91 | (*D*)HB.instance Definition _ := [isSub for eval]. 92 | (*D*)HB.instance Definition _ := [Equality of even_nat by <:]. 93 | 94 | Check forall (n m : even_nat), m == n. 95 | 96 | 97 | 98 | -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise4_solution.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | 3 | Implicit Type P Q R : Prop. 4 | 5 | (** *** Exercise 0: 6 | - Define not. In type theory negation is defined in terms 7 | of [False]. 8 | *) 9 | 10 | Definition not P := 11 | (*D*) P -> False 12 | . 13 | 14 | (** *** Exercise 1: 15 | - Prove the negation of the excluded middle. 16 | *) 17 | Lemma ex0 P : not (P /\ not P). 18 | Proof. 19 | (*D*)by move=> [p np]; apply: np; apply p. 20 | Qed. 21 | 22 | (** *** Exercise 2: 23 | - Declare iff (the constructor being called [iff_intro]). 24 | - Define iff1 of the given type 25 | *) 26 | Inductive iff P Q := 27 | (*D*) iff_intro (p2q : P -> Q) (q2p : Q -> P) 28 | . 29 | 30 | Definition iff1 P Q : iff P Q -> P -> Q := 31 | (*D*) fun x => match x with iff_intro f _ => f end. 32 | 33 | (** *** Exercise 3: 34 | - Declare xor: two constructors, both have two arguments 35 | - Prove the following lemmas 36 | *) 37 | 38 | Inductive xor P Q : Prop := 39 | (*D*) | xorL (p : P) (np : not Q) 40 | (*D*) | xorR (no : not P) (q : Q) 41 | . 42 | 43 | Lemma xorC P Q : iff (xor P Q) (xor Q P). 44 | Proof. 45 | (*D*)apply: iff_intro; case=> [p nq|np q]; by [ apply: xorL | apply: xorR ]. 46 | Qed. 47 | 48 | 49 | Lemma xor1 P Q : (xor P Q) -> not Q -> P. 50 | Proof. 51 | (*D*)case=> [p _| np q] nq. 52 | (*D*) by apply: p. 53 | (*D*)case: (nq q). 54 | Qed. 55 | 56 | Lemma xor2 P Q Z : (xor P Q) -> (xor Q Z) -> iff P Z. 57 | Proof. 58 | (*D*)by case=> [??|??]; case=>[??|??]. 59 | Qed. 60 | 61 | (** *** Exercise 4: 62 | - Declare exists2 63 | - Prove a lemma ex1 -> ex2 T 64 | *) 65 | 66 | Inductive ex2 T (P Q : pred T) : Prop := 67 | (*D*) ex2_intro (x : T) (p : P x) (q : Q x) 68 | . 69 | 70 | Lemma ex2T T (P : pred T) : (exists x, P x) -> (ex2 T P P). 71 | Proof. 72 | (*D*)by case=> x px; apply: (ex2_intro _ _ _ x). 73 | Qed. 74 | 75 | (** *** Exercise 5: 76 | - Write the induction principle for lists 77 | *) 78 | Definition induction_seq A (P : seq A -> Prop) : 79 | P nil -> (forall a l, P l -> P (a :: l)) -> forall l, P l := 80 | (*D*) fun pn pc => 81 | (*D*) fix IH l : P l := 82 | (*D*) match l with nil => pn | cons a l1 => pc a l1 (IH l1) end 83 | . 84 | 85 | 86 | (** *** Exercise 6: 87 | - remeber [=> /view] to prove the following lemma 88 | - the two relevant views are [prime_gt1] and [dvdn_leq] 89 | - Note: [=> /view] combines well with [->] (lesson 3) 90 | - Hint: the proof can be a one liner [by move=> ....] 91 | - Recall: the notation "_ < _ <= _" hides a conjunction 92 | *) 93 | About prime_gt1. 94 | About dvdn_leq. 95 | 96 | Lemma ex_view p : prime p -> p %| 7 -> 1 < p <= 7. 97 | Proof. 98 | (*D*)by move=> /prime_gt1 -> /dvdn_leq ->. 99 | Qed. 100 | 101 | (** *** Exercise 7: 102 | - Define the indexed data type of Cherry tree: 103 | + the index is a bool and must be truee iff the tree is completely 104 | flourished 105 | + leaves can be either Flower or Bud 106 | + the third constructor is called Node and has two sub trees 107 | *) 108 | Inductive cherryt : bool -> Type := 109 | (*D*) | Flower : cherryt true 110 | (*D*) | Bud : cherryt false 111 | (*D*) | Node f (l : cherryt f) (r :cherryt f) : cherryt f. 112 | 113 | Check Node _ Flower Flower : cherryt true. 114 | Check Node _ Bud Bud : cherryt false. 115 | Fail Check Node _ Flower Bud. 116 | 117 | 118 | -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise2_todo.v: -------------------------------------------------------------------------------- 1 | (** Cheat sheet available at 2 | #https://www-sop.inria.fr/teams/marelle/types18/cheatsheet.pdf# 3 | *) 4 | 5 | From mathcomp Require Import all_ssreflect. 6 | 7 | Implicit Type p q r : bool. 8 | Implicit Type m n a b c : nat. 9 | 10 | (** *** 11 | Try to prove the following theorems using no 12 | lemma and minimizing the number of applications of 13 | the tactic case 14 | *) 15 | 16 | (** *** Exercise 1: 17 | *) 18 | 19 | Lemma andTb p : true && p = p. 20 | 21 | (** *** Exercise 2: 22 | *) 23 | 24 | Lemma andbT p : p && true = p. 25 | 26 | (** *** Exercise 3: 27 | *) 28 | 29 | Lemma orbC p q : p || q = q || p. 30 | 31 | (** *** Exercise 4: 32 | *) 33 | Goal forall p q, (p && q) || ( p && ~~ q) || 34 | (~~ p && q) || (~~ p && ~~ q). 35 | 36 | (** *** Exercise 5 : 37 | *) 38 | Goal forall p q r, (p || q) && r = r && (p || q). 39 | 40 | Goal forall n, n < n.+1. 41 | by []. 42 | Qed. 43 | 44 | (** *** Exercise 6 : 45 | - look up what [==>] 46 | *) 47 | Lemma implybE p q : p ==> q = ~~ p || q. 48 | 49 | (** *** Exercise 7 : 50 | Try to prove using the case tactic and alternatively 51 | without using the case tactic 52 | *) 53 | 54 | Lemma negb_imply p q : ~~ (p ==> q) = p && ~~ q. 55 | 56 | 57 | (** *** Exercise 8 : 58 | Try to prove using the case tactic and alternatively 59 | without using the case tactic 60 | *) 61 | Lemma Peirce p q : ((p ==> q) ==> p) ==> p. 62 | 63 | 64 | (** *** Exercise 9 : 65 | - what is [(+)] ? 66 | - prove this using move and rewrite 67 | *) 68 | Lemma find_me p q : ~~ p = q -> p (+) q. 69 | 70 | 71 | (** *** 72 | maxn defines the maximum of two numbers 73 | *) 74 | 75 | Print maxn. 76 | Search maxn in ssrnat. 77 | 78 | (** *** 79 | We define the maxinum of three number as 80 | folllow 81 | *) 82 | 83 | Definition max3n a b c := 84 | if a < b then maxn b c else maxn a c. 85 | 86 | (** *** 87 | Try to prove the following theorem 88 | (you may use properties of maxn) 89 | *) 90 | 91 | 92 | (** *** Exercise 10 93 | *) 94 | 95 | Lemma max3n3n a : max3n a a a = a. 96 | 97 | (** *** Exercise 11 98 | *) 99 | Lemma max3E a b c : max3n a b c = maxn (maxn a b) c. 100 | 101 | (** *** Exercise 12 102 | *) 103 | Lemma max3n_213 a b c : max3n a b c = max3n b a c. 104 | 105 | (** *** Exercise 13 106 | *) 107 | Lemma max3n_132 a b c : max3n a b c = max3n a c b. 108 | 109 | (** *** Exercise 14 110 | *) 111 | Lemma max3n_231 a b c : max3n a b c = max3n b c a. 112 | 113 | (** *** 114 | We define functions that test if 3 natural numbers are 115 | in increasing (or decreasing) order 116 | *) 117 | 118 | Definition order3n (T : Type) (r : rel T) x y z := (r x y) && (r y z). 119 | Definition incr3n := order3n nat (fun x y => x <= y). 120 | Definition decr3n := order3n nat (fun x y => y <= x). 121 | 122 | (** *** Exercise 15 123 | *) 124 | Lemma incr3n_decr a b c : incr3n a b c = decr3n c b a. 125 | 126 | (** *** Exercise 16 127 | *) 128 | 129 | Lemma incr3_3n a : incr3n a a a. 130 | 131 | (** *** Exercise 17 132 | *) 133 | 134 | Lemma decr3_3n a : decr3n a a a. 135 | 136 | (** *** Exercise 18 137 | *) 138 | 139 | Lemma incr3n_leq12 a b c : incr3n a b c -> a <= b. 140 | 141 | (** *** Exercise 19 142 | *) 143 | Lemma incr3n_leq23 a b c : incr3n a b c -> b <= c. 144 | 145 | (** *** Exercise 20 146 | *) 147 | Lemma incr3n_eq a b c : incr3n a b a = (a == b). 148 | 149 | -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise2_solution.v: -------------------------------------------------------------------------------- 1 | (** Cheat sheet available at 2 | #https://www-sop.inria.fr/teams/marelle/types18/cheatsheet.pdf# 3 | *) 4 | 5 | From mathcomp Require Import all_ssreflect. 6 | 7 | Implicit Type p q r : bool. 8 | Implicit Type m n a b c : nat. 9 | 10 | (** *** 11 | Try to prove the following theorems using no 12 | lemma and minimizing the number of applications of 13 | the tactic case 14 | *) 15 | 16 | (** *** Exercise 1: 17 | *) 18 | 19 | Lemma andTb p : true && p = p. 20 | (*D*)Proof. by []. Qed. 21 | 22 | (** *** Exercise 2: 23 | *) 24 | 25 | Lemma andbT p : p && true = p. 26 | (*D*)Proof. by case: p. Qed. 27 | 28 | (** *** Exercise 3: 29 | *) 30 | 31 | Lemma orbC p q : p || q = q || p. 32 | (*D*)Proof. by case: p; case: q. Qed. 33 | 34 | (** *** Exercise 4: 35 | *) 36 | Goal forall p q, (p && q) || ( p && ~~ q) || 37 | (~~ p && q) || (~~ p && ~~ q). 38 | (*D*)Proof. by move=> p q; case: p; case: q. Qed. 39 | 40 | (** *** Exercise 5 : 41 | *) 42 | Goal forall p q r, (p || q) && r = r && (p || q). 43 | (*D*)Proof. by move=> p q r; case: (p || q); case: r. Qed. 44 | 45 | Goal forall n, n < n.+1. 46 | by []. 47 | Qed. 48 | 49 | (** *** Exercise 6 : 50 | - look up what [==>] 51 | *) 52 | (*D*)Locate "==>". 53 | (*D*)Print implb. 54 | Lemma implybE p q : p ==> q = ~~ p || q. 55 | (*D*) Proof. by case: p. Qed. 56 | 57 | (** *** Exercise 7 : 58 | Try to prove using the case tactic and alternatively 59 | without using the case tactic 60 | *) 61 | 62 | Lemma negb_imply p q : ~~ (p ==> q) = p && ~~ q. 63 | (*D*) (* Proof. by case: p. Qed. *) 64 | (*D*) Proof. by rewrite implybE negb_or negbK. Qed. 65 | 66 | 67 | (** *** Exercise 8 : 68 | Try to prove using the case tactic and alternatively 69 | without using the case tactic 70 | *) 71 | Lemma Peirce p q : ((p ==> q) ==> p) ==> p. 72 | (*D*) (* Proof. by case: p; case: q. Qed. *) 73 | (*D*) Proof. by rewrite implybE negb_imply implybE orbK orNb. Qed. 74 | 75 | 76 | (** *** Exercise 9 : 77 | - what is [(+)] ? 78 | - prove this using move and rewrite 79 | *) 80 | Lemma find_me p q : ~~ p = q -> p (+) q. 81 | (*D*)Locate "(+)". 82 | (*D*)Search _ addb negb. 83 | (*D*)Proof. by move=> np_q; rewrite -np_q addbN negb_add. Qed. 84 | 85 | 86 | (** *** 87 | maxn defines the maximum of two numbers 88 | *) 89 | 90 | Print maxn. 91 | Search maxn inside ssrnat. 92 | 93 | (** *** 94 | We define the maxinum of three number as 95 | folllow 96 | *) 97 | 98 | Definition max3n a b c := 99 | if a < b then maxn b c else maxn a c. 100 | 101 | (** *** 102 | Try to prove the following theorem 103 | (you may use properties of maxn) 104 | *) 105 | 106 | 107 | (** *** Exercise 10 108 | *) 109 | 110 | Lemma max3n3n a : max3n a a a = a. 111 | (*D*) Proof. by rewrite /max3n if_same maxnn. Qed. 112 | 113 | (** *** Exercise 11 114 | *) 115 | Lemma max3E a b c : max3n a b c = maxn (maxn a b) c. 116 | (*D*) Proof. by rewrite /max3n /maxn; case: (a < b). Qed. 117 | 118 | (** *** Exercise 12 119 | *) 120 | Lemma max3n_213 a b c : max3n a b c = max3n b a c. 121 | (*D*) Proof. by rewrite max3E (maxnC a) -max3E. Qed. 122 | 123 | (** *** Exercise 13 124 | *) 125 | Lemma max3n_132 a b c : max3n a b c = max3n a c b. 126 | (*D*) Proof. by rewrite max3E -maxnA (maxnC b) maxnA -max3E. Qed. 127 | 128 | (** *** Exercise 14 129 | *) 130 | Lemma max3n_231 a b c : max3n a b c = max3n b c a. 131 | (*D*) Proof. by rewrite max3n_213 max3n_132. Qed. 132 | 133 | (** *** 134 | We define functions that test if 3 natural numbers are 135 | in increasing (or decreasing) order 136 | *) 137 | 138 | Definition order3n (T : Type) (r : rel T) x y z := (r x y) && (r y z). 139 | Definition incr3n := order3n nat (fun x y => x <= y). 140 | Definition decr3n := order3n nat (fun x y => y <= x). 141 | 142 | (** *** Exercise 15 143 | *) 144 | Lemma incr3n_decr a b c : incr3n a b c = decr3n c b a. 145 | (*D*) Proof. by rewrite /incr3n /order3n andbC. Qed. 146 | 147 | (** *** Exercise 16 148 | *) 149 | 150 | Lemma incr3_3n a : incr3n a a a. 151 | (*D*) by rewrite /incr3n /order3n leqnn. Qed. 152 | 153 | (** *** Exercise 17 154 | *) 155 | 156 | Lemma decr3_3n a : decr3n a a a. 157 | (*D*) by rewrite -incr3n_decr incr3_3n. Qed. 158 | 159 | (** *** Exercise 18 160 | *) 161 | 162 | Lemma incr3n_leq12 a b c : incr3n a b c -> a <= b. 163 | (*D*) by rewrite /incr3n /order3n; case: (_ <= _). Qed. 164 | 165 | (** *** Exercise 19 166 | *) 167 | Lemma incr3n_leq23 a b c : incr3n a b c -> b <= c. 168 | (*D*) by rewrite /incr3n /order3n; case: (_ <= _). Qed. 169 | 170 | (** *** Exercise 20 171 | *) 172 | Lemma incr3n_eq a b c : incr3n a b a = (a == b). 173 | (*D*) by rewrite /incr3n /order3n eqn_leq. Qed. 174 | 175 | -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise3_todo.v: -------------------------------------------------------------------------------- 1 | (** Cheat sheet available at 2 | #https://www-sop.inria.fr/teams/marelle/types18/cheatsheet.pdf# 3 | *) 4 | 5 | From mathcomp Require Import all_ssreflect. 6 | 7 | Implicit Type p q r : bool. 8 | Implicit Type m n a b c : nat. 9 | 10 | (** 11 | ----------------------------------------------------------------- 12 | #
# 13 | *** Exercise 1: 14 | - prove this satement by induction 15 | #
# 16 | *) 17 | Lemma iterSr A n (f : A -> A) x : iter n.+1 f x = iter n f (f x). 18 | Admitted. 19 | 20 | (** 21 | #
# 22 | 23 | #


# 24 | #

# 25 | 26 | ---------------------------------------------------------- 27 | #
# 28 | 29 | *** Exercise 2: 30 | - look up the definition of [iter] (note there is an accumulator varying 31 | during recursion) 32 | - prove the following statement by induction 33 | 34 | #
# 35 | *) 36 | Lemma iter_predn m n : iter n predn m = m - n. 37 | Proof. 38 | Admitted. 39 | (** 40 | #
# 41 | 42 | #


# 43 | #

# 44 | 45 | ---------------------------------------------------------- 46 | #
# 47 | 48 | *** Exercise 3: 49 | 50 | Prove the sum of the lists [odds n] of exercise 1 is [n ^ 2]. 51 | 52 | You can prove the following lemmas in any order, some are way easier 53 | than others. 54 | 55 | - Recall from exercise 1. 56 | #
# 57 | 58 | *) 59 | Definition add2list s := map (fun x => x.+2) s. 60 | Definition odds n := iter n (fun s => 1 :: add2list s) [::]. 61 | (** 62 | #
# 63 | 64 | - We define a sum operation [suml]. 65 | 66 | #
# 67 | 68 | *) 69 | Definition suml := foldl addn 0. 70 | (** 71 | #
# 72 | 73 | - Any [foldl addn] can be rexpressed as a sum. 74 | 75 | #
# 76 | *) 77 | Lemma foldl_addE n s : foldl addn n s = n + suml s. 78 | Proof. 79 | Admitted. 80 | (** 81 | #
# 82 | 83 | 84 | - Not to break abstraction, prove [suml_cons]. 85 | #
# 86 | *) 87 | Lemma suml_cons n s : suml (n :: s) = n + suml s. 88 | Admitted. 89 | (** 90 | #
# 91 | 92 | - Show how to sum a [add2list]. 93 | 94 | #
# 95 | *) 96 | Lemma suml_add2list s : suml (add2list s) = suml s + 2 * size s. 97 | Proof. 98 | Admitted. 99 | (** 100 | #
# 101 | 102 | - Show the size of a [add2list]. 103 | 104 | #
# 105 | *) 106 | Lemma size_add2list s : size (add2list s) = size s. 107 | Admitted. 108 | (** 109 | #
# 110 | 111 | - Show how many elments [odds] have. 112 | 113 | #
# 114 | *) 115 | Lemma size_odds n : size (odds n) = n. 116 | Admitted. 117 | (** 118 | #
# 119 | - Show the final statment. 120 | #
# 121 | *) 122 | Lemma eq_suml_odds n : suml (odds n) = n ^ 2. 123 | Proof. 124 | Admitted. 125 | (** 126 | #
# 127 | #
# 128 | #
(hint)
# 129 | For [eq_suml_odds], use [sqrnD] 130 | #
# 131 | #


# 132 | #

# 133 | 134 | ---------------------------------------------------------- 135 | #
# 136 | 137 | *** Exercise 4: 138 | 139 | Prove the sum of the lists [odds n] is what you think. 140 | 141 | You may want to have at least one itermediate lemma to prove [oddsE] 142 | #
# 143 | *) 144 | Lemma oddsE n : odds n = [seq 2 * i + 1 | i <- iota 0 n]. 145 | Proof. 146 | Admitted. 147 | (** 148 | #
# 149 | #
# 150 | #
(hint)
# 151 | this intermediate lemma would be: 152 | #
# 153 | *) 154 | Lemma oddsE_aux n k : 155 | iter n (fun s : seq nat => 2 * k + 1 :: add2list s) [::] = 156 | [seq 2 * i + 1 | i <- iota k n]. 157 | Admitted. 158 | (** 159 | #
# 160 | #
# 161 | #
# 162 | #
# 163 | *) 164 | Lemma nth_odds n i : i < n -> nth 0 (odds n) i = 2 * i + 1. 165 | Proof. 166 | Admitted. 167 | (** 168 | #
# 169 | #


# 170 | #

# 171 | 172 | ---------------------------------------------------------- 173 | #
# 174 | 175 | *** Exercise 5: 176 | 177 | Let us prove directly formula 178 | #$$ \sum_{i=0}^{n-1} (2 i + 1) = n ^ 2 $$# 179 | from lesson 1, slightly modified. 180 | 181 | Let us first define a custom sum operator: 182 | #
# 183 | *) 184 | Definition mysum m n F := (foldr (fun i a => F i + a) 0 (iota m (n - m))). 185 | 186 | Notation "\mysum_ ( m <= i < n ) F" := (mysum m n (fun i => F)) 187 | (at level 41, F at level 41, i, m, n at level 50, 188 | format "'[' \mysum_ ( m <= i < n ) '/ ' F ']'"). 189 | (** 190 | #
# 191 | - First prove a very useful lemma about summation 192 | #
# 193 | *) 194 | Lemma mysum_recl m n F : m <= n -> 195 | \mysum_(m <= i < n.+1) F i = \mysum_(m <= i < n) F i + F n. 196 | Proof. 197 | Admitted. 198 | (** 199 | #
# 200 | - Now prove the main result 201 | 202 | Do NOT use [eq_suml_odds] above, it would take much more time 203 | 204 | #
# 205 | *) 206 | Lemma sum_odds n : \mysum_(0 <= i < n) (2 * i + 1) = n ^ 2. 207 | Proof. 208 | Admitted. 209 | (** 210 | #
# 211 | #


# 212 | #

# 213 | *) -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise5_todo.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | 3 | Module easy. 4 | 5 | Implicit Type p q r : bool. 6 | Implicit Type m n a b c : nat. 7 | 8 | (** *** Exercise 1: 9 | - look for lemmas supporting contrapositive reasoning 10 | - use the eqP view to finish the proof. 11 | *) 12 | Lemma bool_gimmics1 a : a != a.-1 -> a != 0. 13 | 14 | (** *** Exercise 2: 15 | - it helps to find out what is behind [./2] and [.*2] in order to [Search] 16 | - any proof would do, but there is one not using [implyP] 17 | *) 18 | Lemma view_gimmics1 p a b : p -> (p ==> (a == b.*2)) -> a./2 = b. 19 | 20 | (** *** Exercise 3: 21 | - Prove this view by unfolding maxn and then using [leqP] 22 | *) 23 | Lemma maxn_idPl m n : reflect (maxn m n = m) (m >= n). 24 | Proof. apply: (iffP idP). 25 | Qed. 26 | 27 | (** *** Exercise 4: 28 | - there is no need to prove [reflect] with [iffP]: here just use [rewrite] and [apply] 29 | - check out the definitions and theory of [leq] and [maxn] 30 | - proof sketch: 31 | << 32 | n <= m = n - m == 0 33 | = m + n - m == m + 0 34 | = maxn m n == m 35 | >> *) 36 | Lemma maxn_idPl_bis m n : reflect (maxn m n = m) (m >= n). 37 | 38 | End easy. 39 | 40 | 41 | 42 | Module reflect1. 43 | 44 | 45 | (** *** Exercise 5: 46 | - Prove some reflection lemmas for the proposed reflect definition 47 | *) 48 | 49 | (** 50 | - A possible definition for reflect 51 | *) 52 | 53 | Inductive reflectl (P : Prop) b := 54 | | RT (p : P) (e : b = true) 55 | | RF (p : ~ P) (e : b = false). 56 | About reflect. 57 | 58 | 59 | Lemma andP (b1 b2 : bool) : reflectl (b1 /\ b2) (b1 && b2). 60 | Proof. 61 | Qed. 62 | 63 | Lemma orP (b1 b2 : bool) : reflectl (b1 \/ b2) (b1 || b2). 64 | Proof. 65 | Qed. 66 | 67 | Lemma implyP (b1 b2 : bool) : reflectl (b1 -> b2) (b1 ==> b2). 68 | Proof. 69 | Qed. 70 | 71 | End reflect1. 72 | 73 | 74 | 75 | 76 | 77 | (** *** Exercise 6: 78 | - Get excluded-middle when P is equivalent to a "bool" "decidable" 79 | *) 80 | 81 | (** Equivalence definition *) 82 | 83 | Definition bool_Prop_equiv (P : Prop) (b : bool) := b = true <-> P. 84 | Lemma test_bool_Prop_equiv b P : bool_Prop_equiv P b -> P \/ ~ P. 85 | Proof. 86 | Qed. 87 | 88 | (** *** Exercise 7: 89 | - Let's use standard reflect (and iffP, idP etc...) 90 | *) 91 | Lemma iffP_lr (P : Prop) (b : bool) : 92 | (P -> b) -> (b -> P) -> reflect P b. 93 | Proof. 94 | Qed. 95 | 96 | Lemma iffP_rl (P : Prop) (b : bool) : 97 | reflect P b -> ((P -> b) /\ (b -> P)). 98 | Proof. 99 | Qed. 100 | 101 | 102 | Lemma eqnP {n m : nat} : 103 | reflect (n = m) (eqn n m). 104 | Proof. 105 | Qed. 106 | 107 | 108 | (** *** Exercise 8: 109 | If you have time.. more reflections 110 | 111 | - leq_max : use leq_total maxn_idPl 112 | - dvdn_fact: use leq_eqVlt dvdn_mulr dvdn_mull 113 | *) 114 | 115 | Lemma nat_inj_eq T (f : T -> nat) x y : 116 | injective f -> 117 | reflect (x = y) (eqn (f x) (f y)). 118 | Proof. 119 | Qed. 120 | 121 | Lemma leq_max m n1 n2 : (m <= maxn n1 n2) = (m <= n1) || (m <= n2). 122 | Proof. 123 | Qed. 124 | 125 | 126 | Lemma dvdn_fact m n : 0 < m <= n -> m %| n`!. 127 | Proof. 128 | Qed. 129 | 130 | Lemma prime_above m : exists2 p, m < p & prime p. 131 | Proof. 132 | Qed. 133 | 134 | (** 135 | For this section: 136 | only move=> h, move/V: h => h, case/V: h, by ... allowed 137 | *) 138 | 139 | Goal forall (P Q : Prop), (P <-> Q) -> P -> Q. 140 | Proof. 141 | Qed. 142 | 143 | Goal forall (P : nat -> Prop) (Q : Prop), 144 | (P 0 -> Q) 145 | -> (forall n, P n.+1 -> P n) 146 | -> P 4 -> Q. 147 | Proof. 148 | Qed. 149 | 150 | 151 | (** No case analysis on b, b1, b2 allowed *) 152 | Goal forall (b b1 b2 : bool), (b1 -> b) -> (b2 -> b) -> b1 || b2 -> b. 153 | Proof. 154 | Qed. 155 | 156 | Goal forall (Q : nat -> Prop) (p1 p2 : nat -> bool) x, 157 | ((forall y, Q y -> p1 y /\ p2 y) /\ Q x) -> p1 x && p2 x. 158 | Proof. 159 | Qed. 160 | 161 | Goal forall (Q : nat -> Prop) (p1 p2 : nat -> bool) x, 162 | ((forall y, Q y -> p1 y \/ p2 y) /\ Q x) -> p1 x || p2 x. 163 | Proof. 164 | Qed. 165 | 166 | (** 167 | No xxxP lemmas allowed, but reflectT and reflectF and case analysis allowed , 168 | *) 169 | Lemma myidP: forall (b : bool), reflect b b. 170 | Proof. 171 | Qed. 172 | 173 | Lemma mynegP: forall (b : bool), reflect (~ b) (~~ b). 174 | Proof. 175 | Qed. 176 | 177 | Lemma myandP: forall (b1 b2 : bool), reflect (b1 /\ b2) (b1 && b2). 178 | Proof. 179 | Qed. 180 | 181 | Lemma myiffP: 182 | forall (P Q : Prop) (b : bool), 183 | reflect P b -> (P -> Q) -> (Q -> P) -> reflect Q b. 184 | Proof. 185 | Qed. 186 | 187 | 188 | (** 189 | Some arithmetics 190 | *) 191 | 192 | Fixpoint len (n m : nat) : bool := 193 | match n, m with 194 | | 0 , _ => true 195 | | n'.+1, 0 => false 196 | | n'.+1, m'.+1 => len n' m' 197 | end. 198 | 199 | Lemma lenP: forall n m, reflect (exists k, k + n = m) (len n m). 200 | Proof. 201 | Qed. 202 | 203 | 204 | (* --------------------------------------------------------------------*) 205 | -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise7_todo.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | 3 | 4 | (** *** Exercise 1 : 5 | - Prove this statement by induction or 6 | - alternatively by using big_morph 7 | *) 8 | 9 | Lemma sum_mull f (k n : nat) : 10 | k * (\sum_(0 <= i < n) f i) = \sum_(0 <= i < n) (k * f i). 11 | Proof. 12 | Qed. 13 | 14 | (** *** Exercise 2 : 15 | - Prove this statement by using big_morph 16 | *) 17 | 18 | Lemma sum_mulr f (k n : nat) : 19 | (\sum_(0 <= i < n) f i) * k = \sum_(0 <= i < n) (f i * k). 20 | Proof. 21 | Qed. 22 | 23 | 24 | (** *** Exercise 3 : 25 | - Prove this statement by induction. 26 | - Relevant lemmas are 27 | - doubleS odd_double addn0 addnn mulnn 28 | - big_mkcond big_nat_recr big_geq 29 | *) 30 | 31 | Lemma sum_odd n : \sum_(0 <= i < n.*2 | odd i) i = n ^ 2. 32 | Proof. 33 | Qed. 34 | 35 | 36 | (** *** Exercise 4 : 37 | - Prove this statement by induction. 38 | - Relevant lemmas are 39 | - doubleD muln2 addn2 big_nat_recr big_geq 40 | *) 41 | 42 | Lemma sum_gauss n : (\sum_(0 <= i < n) i).*2 = n * n.-1. 43 | Proof. 44 | Qed. 45 | 46 | 47 | (** 48 | In what follows we are going to mimic the proof of Gauss : 49 | 50 | << 51 | 1 + 2 + ..... + n.-2 + n.-1 52 | + n.-1 + n.-2 + + 2 + 1 53 | 54 | = n.-1 * n 55 | >> 56 | 57 | **) 58 | 59 | 60 | (** *** Exercise 5 : 61 | - Prove this statement without induction. 62 | - Relevant lemma is big_nat_rev 63 | *) 64 | 65 | Lemma sum_gauss_rev n : \sum_(0 <= i < n) (n.-1 - i) = \sum_(0 <= i < n) i. 66 | Proof. 67 | Qed. 68 | 69 | (** *** Exercise 6 : 70 | - Prove this statement without induction. 71 | - Relevant lemma is addnn 72 | *) 73 | Lemma sum_gauss_double n : (\sum_(0 <= i < n) i).*2 = 74 | \sum_(0 <= i < n) i + \sum_(0 <= i < n) (n.-1 - i). 75 | Proof. 76 | Qed. 77 | 78 | 79 | (** *** Exercise 7 : 80 | - Prove this statement without induction. 81 | - Relevant lemma are big_split and eq_big_nat 82 | *) 83 | 84 | Lemma sum_gaussD n : 85 | \sum_(0 <= i < n) i + \sum_(0 <= i < n) (n.-1 - i) = 86 | \sum_(0 <= i < n) n.-1. 87 | Proof. 88 | Qed. 89 | 90 | 91 | (** *** Exercise 8 : 92 | - Prove this statement without induction. 93 | - Relevant lemma are sum_nat_const_nat 94 | *) 95 | 96 | Lemma sum_gauss_const n k : \sum_(0 <= i < n) k = n * k. 97 | Proof. 98 | Qed. 99 | 100 | 101 | (** *** Exercise 9 : 102 | - Prove this statement using exercises 5-8 103 | *) 104 | Lemma sum_gauss_alt1 n : (\sum_(0 <= i < n) i).*2 = n * n.-1. 105 | Proof. 106 | Qed. 107 | 108 | 109 | (** *** Exercise 10 : 110 | - Prove this statement directly without using the lemmas 111 | - defined in exercises 5-9 112 | *) 113 | 114 | Lemma sum_gauss_alt2 n : (\sum_(0 <= i < n) i).*2 = n * n.-1. 115 | Proof. 116 | Qed. 117 | 118 | 119 | (** *** Now we try to prove the sum of squares. 120 | 121 | **) 122 | 123 | (** *** We first define the property for a function to be increasing 124 | **) 125 | 126 | 127 | Definition fincr f := forall n, f n <= f n.+1. 128 | 129 | (** *** Exercise 11 : 130 | - Prove this statement by induction 131 | *) 132 | 133 | Lemma fincrD f m n : fincr f -> f m <= f (n + m). 134 | Proof. 135 | Qed. 136 | 137 | 138 | (** *** Exercise 12 : 139 | - Prove this statement using exercise 9 140 | - Hint : subnK 141 | *) 142 | 143 | Lemma fincr_leq f m n : fincr f -> m <= n -> f m <= f n. 144 | Proof. 145 | Qed. 146 | 147 | 148 | (** *** Exercise 13 : 149 | - Proof by induction 150 | - Hints : addnCA subnK fincr_leq big_geq 151 | *) 152 | 153 | Lemma sum_consecutive n f : 154 | fincr f -> f n = \sum_(0 <= i < n) (f i.+1 - f i) + f 0. 155 | Proof. 156 | Qed. 157 | 158 | 159 | (** *** Exercise 14 : 160 | - Proof using the previous lemma 161 | - Hints : leq_exp2r 162 | *) 163 | Lemma sum_consecutive_cube n : 164 | n^3 = \sum_(0 <= i < n) (i.+1 ^ 3 - i ^ 3). 165 | Proof. 166 | Qed. 167 | 168 | 169 | (** *** We give the proof of a technical result 170 | *) 171 | 172 | Ltac sring := 173 | rewrite !(expn1, expnS, =^~mul2n, mulSn, mulnS, addnS, addSn, 174 | mulnDr, mulnDl, add0n, addn0, muln0, addnA, mulnA); 175 | do ! congr (S _); 176 | do ! ((congr (_ + _); [idtac]) || (rewrite [in LHS]addnC ?[in LHS]addnA //)). 177 | 178 | Lemma succ_cube n : n.+1 ^ 3 = n ^ 3 + (3 * n ^ 2 + 3 * n + 1). 179 | Proof. sring. Qed. 180 | 181 | (** *** Exercise 15 : 182 | - Hints : big_split sum_mll sum_gauss sum_gauss_const 183 | *) 184 | Lemma sum_sum3 n : 185 | \sum_(0 <= i < n) (6 * i ^ 2 + 6 * i + 2) = 186 | 6 * (\sum_(0 <= i < n) i ^ 2) + 3 * n * n.-1 + n.*2. 187 | Proof. 188 | Qed. 189 | 190 | 191 | (** *** Exercise 16 : 192 | - Hints : big_split sum_mll sum_gauss sum_gauss_const 193 | *) 194 | Lemma sum_sum4 n : 195 | (n ^ 3).*2 = 6 * (\sum_(0 <= i < n) i ^ 2) + 3 * n * n.-1 + n.*2. 196 | Proof. 197 | Qed. 198 | 199 | (** *** Another technical lemma 200 | *) 201 | 202 | Lemma sum_tech n : (n ^ 3).*2 = n * n.-1 * n.-1.*2.+1 + 3 * n * n.-1 + n.*2. 203 | Proof. by case: n => //= n1; sring. Qed. 204 | 205 | 206 | (** *** Exercise 17 : 207 | - Hint : addIn sum_sum4 sum_tech. 208 | *) 209 | Lemma sum_square n : 6 * (\sum_(0 <= i < n) i ^ 2) = n * n.-1 * n.-1.*2.+1. 210 | Proof. 211 | Qed. 212 | 213 | (** *** Exercise 18 : 214 | - Prove the theorem directly using only sum_gauss and the tactic sring. 215 | *) 216 | Lemma sum_square_alt n : 6 * (\sum_(0 <= i < n) i ^ 2) = n * n.-1 * n.-1.*2.+1. 217 | Proof. 218 | Qed. 219 | -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise3_solution.v: -------------------------------------------------------------------------------- 1 | (** Cheat sheet available at 2 | #https://www-sop.inria.fr/teams/marelle/types18/cheatsheet.pdf# 3 | *) 4 | 5 | From mathcomp Require Import all_ssreflect. 6 | 7 | Implicit Type p q r : bool. 8 | Implicit Type m n a b c : nat. 9 | 10 | (** 11 | ----------------------------------------------------------------- 12 | #
# 13 | *** Exercise 1: 14 | - prove this satement by induction 15 | #
# 16 | *) 17 | Lemma iterSr A n (f : A -> A) x : iter n.+1 f x = iter n f (f x). 18 | (*A*)Proof. by elim: n => //= n <-. Qed. 19 | 20 | (** 21 | #
# 22 | 23 | #


# 24 | #

# 25 | 26 | ---------------------------------------------------------- 27 | #
# 28 | 29 | *** Exercise 2: 30 | - look up the definition of [iter] (note there is an accumulator varying 31 | during recursion) 32 | - prove the following statement by induction 33 | 34 | #
# 35 | *) 36 | Lemma iter_predn m n : iter n predn m = m - n. 37 | Proof. 38 | (*D*)by elim: n m => [|n IHn] m; rewrite ?subn0//= IHn subnS. 39 | (*A*)Qed. 40 | (** 41 | #
# 42 | 43 | #


# 44 | #

# 45 | 46 | ---------------------------------------------------------- 47 | #
# 48 | 49 | *** Exercise 3: 50 | 51 | Prove the sum of the lists [odds n] of exercise 1 is [n ^ 2]. 52 | 53 | You can prove the following lemmas in any order, some are way easier 54 | than others. 55 | 56 | - Recall from exercise 1. 57 | #
# 58 | 59 | *) 60 | Definition add2list s := map (fun x => x.+2) s. 61 | Definition odds n := iter n (fun s => 1 :: add2list s) [::]. 62 | (** 63 | #
# 64 | 65 | - We define a sum operation [suml]. 66 | 67 | #
# 68 | 69 | *) 70 | Definition suml := foldl addn 0. 71 | (** 72 | #
# 73 | 74 | - Any [foldl addn] can be rexpressed as a sum. 75 | 76 | #
# 77 | *) 78 | Lemma foldl_addE n s : foldl addn n s = n + suml s. 79 | Proof. 80 | (*D*)elim: s n => [n | x s IHs n] /=; first by rewrite addn0. 81 | (*D*)by rewrite /suml/= !IHs add0n addnA. 82 | (*A*)Qed. 83 | (** 84 | #
# 85 | 86 | 87 | - Not to break abstraction, prove [suml_cons]. 88 | #
# 89 | *) 90 | Lemma suml_cons n s : suml (n :: s) = n + suml s. 91 | (*A*)Proof. by rewrite /suml/= foldl_addE. Qed. 92 | (** 93 | #
# 94 | 95 | - Show how to sum a [add2list]. 96 | 97 | #
# 98 | *) 99 | Lemma suml_add2list s : suml (add2list s) = suml s + 2 * size s. 100 | Proof. 101 | (*D*)elim: s => [|x s IHs] //=; rewrite !suml_cons IHs. 102 | (*D*)by rewrite !mulnS !addnS addnA. 103 | (*A*)Qed. 104 | (** 105 | #
# 106 | 107 | - Show the size of a [add2list]. 108 | 109 | #
# 110 | *) 111 | Lemma size_add2list s : size (add2list s) = size s. 112 | (*A*)Proof. by apply: size_map. Qed. 113 | (** 114 | #
# 115 | 116 | - Show how many elments [odds] have. 117 | 118 | #
# 119 | *) 120 | Lemma size_odds n : size (odds n) = n. 121 | (*A*)Proof. by elim: n => //= n; rewrite size_add2list => ->. Qed. 122 | (** 123 | #
# 124 | - Show the final statment. 125 | #
# 126 | *) 127 | Lemma eq_suml_odds n : suml (odds n) = n ^ 2. 128 | Proof. 129 | (*D*)elim: n => //= n IHn; rewrite suml_cons. 130 | (*D*)rewrite suml_add2list IHn addnCA addnA. 131 | (*D*)by rewrite -(addn1 n) sqrnD size_odds muln1. 132 | (*A*)Qed. 133 | (** 134 | #
# 135 | #
# 136 | #
(hint)
# 137 | For [eq_suml_odds], use [sqrnD] 138 | #
# 139 | #


# 140 | #

# 141 | 142 | ---------------------------------------------------------- 143 | #
# 144 | 145 | *** Exercise 4: 146 | 147 | Prove the sum of the lists [odds n] is what you think. 148 | 149 | You may want to have at least one itermediate lemma to prove [oddsE] 150 | #
# 151 | *) 152 | Lemma oddsE n : odds n = [seq 2 * i + 1 | i <- iota 0 n]. 153 | Proof. 154 | (*D*)rewrite /odds; set k := (0 as X in iota X _). 155 | (*D*)rewrite -[1 in LHS]/(2 * k + 1). 156 | (*D*)elim: n k => //= n IHn k; rewrite {}IHn; congr (_ :: _). 157 | (*D*)by elim: n k => // n IHn k /=; rewrite IHn !mulnS. 158 | (*A*)Qed. 159 | (** 160 | #
# 161 | #
# 162 | #
(hint)
# 163 | this intermediate lemma would be: 164 | #
# 165 | *) 166 | Lemma oddsE_aux n k : 167 | iter n (fun s : seq nat => 2 * k + 1 :: add2list s) [::] = 168 | [seq 2 * i + 1 | i <- iota k n]. 169 | (*D*)Proof. 170 | (*D*)elim: n k => //= n IHn k; rewrite {}IHn; congr (_ :: _). 171 | (*D*)by elim: n k => // n IHn k /=; rewrite IHn !mulnS. 172 | (*A*)Qed. 173 | (** 174 | #
# 175 | #
# 176 | #
# 177 | #
# 178 | *) 179 | Lemma nth_odds n i : i < n -> nth 0 (odds n) i = 2 * i + 1. 180 | Proof. 181 | (*D*)by move=> i_ltn; rewrite oddsE (nth_map 0) ?size_iota ?nth_iota. 182 | (*A*)Qed. 183 | (** 184 | #
# 185 | #


# 186 | #

# 187 | 188 | ---------------------------------------------------------- 189 | #
# 190 | 191 | *** Exercise 5: 192 | 193 | Let us prove directly formula 194 | #$$ \sum_{i=0}^{n-1} (2 i + 1) = n ^ 2 $$# 195 | from lesson 1, slightly modified. 196 | 197 | Let us first define a custom sum operator: 198 | #
# 199 | *) 200 | Definition mysum m n F := (foldr (fun i a => F i + a) 0 (iota m (n - m))). 201 | 202 | Notation "\mysum_ ( m <= i < n ) F" := (mysum m n (fun i => F)) 203 | (at level 41, F at level 41, i, m, n at level 50, 204 | format "'[' \mysum_ ( m <= i < n ) '/ ' F ']'"). 205 | (** 206 | #
# 207 | - First prove a very useful lemma about summation 208 | #
# 209 | *) 210 | Lemma mysum_recl m n F : m <= n -> 211 | \mysum_(m <= i < n.+1) F i = \mysum_(m <= i < n) F i + F n. 212 | Proof. 213 | (*D*)move=> leq_mn; rewrite /mysum subSn// -addn1 iotaD subnKC//= foldr_cat/=. 214 | (*D*)elim: (iota _ _) (F n) => [|x s IHs] k /=; first by rewrite addn0. 215 | (*D*)by rewrite IHs addnA. 216 | (*A*)Qed. 217 | (** 218 | #
# 219 | - Now prove the main result 220 | 221 | Do NOT use [eq_suml_odds] above, it would take much more time 222 | 223 | #
# 224 | *) 225 | Lemma sum_odds n : \mysum_(0 <= i < n) (2 * i + 1) = n ^ 2. 226 | Proof. 227 | (*D*)elim: n => // n IHn; rewrite mysum_recl// IHn. 228 | (*D*)by rewrite -[n.+1]addn1 sqrnD muln1 addnAC addnA. 229 | (*A*)Qed. 230 | (** 231 | #
# 232 | #


# 233 | #

# 234 | *) -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise6_solution.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | Set Implicit Arguments. 3 | Unset Strict Implicit. 4 | Unset Printing Implicit Defensive. 5 | 6 | Implicit Type p q r : bool. 7 | Implicit Type m n a b c : nat. 8 | 9 | (** *** Exercise 1: 10 | 11 | Prove that the equation #$$ 8y = 6x + 1 $$# has no solution. 12 | 13 | - Hint 1: take the modulo 2 of the equation. 14 | - Hint 2: [Search _ modn addn] and [Search _ modn muln] 15 | #
# *) 16 | Lemma ex1 x y : 8 * y != 6 * x + 1. 17 | Proof. 18 | (*D*)apply/negP => /eqP /(congr1 (modn^~ 2)). 19 | (*D*)by rewrite -modnMml mul0n -modnDml -modnMml. 20 | (*A*)Qed. 21 | 22 | (** #
# 23 | *** Exercise 2: 24 | 25 | The ultimate Goal of this exercise is to find the solutions of the equation 26 | #$$ 2^n = a^2 + b^2,$$# where n is fixed and a and b unkwown. 27 | 28 | We hence study the following predicate: 29 | #
# *) 30 | Definition sol n a b := [&& a > 0, b > 0 & 2 ^ n == a ^ 2 + b ^ 2]. 31 | (** #
# 32 | - First prove that there are no solutions when n is 0. 33 | 34 | - Hint: do enough cases on a and b. 35 | #
# *) 36 | Lemma sol0 a b : ~~ sol 0 a b. 37 | (*A*)Proof. by move: a b => [|[|[|a]]] [|[|[|b]]]. Qed. 38 | (** #
# 39 | - Now prove the only solution when n is 1. 40 | 41 | - Hint: do enough cases on a and b. 42 | #
# *) 43 | Lemma sol1 a b : sol 1 a b = (a == 1) && (b == 1). 44 | (*A*)Proof. by move: a b => [|[|[|a]]] [|[|[|b]]]. Qed. 45 | (** #
# 46 | - Now prove a little lemma that will guarantee that a and b are even. 47 | 48 | - Hint 1: first prove [(x * 2 + y) ^ 2 = y ^ 2 %[mod 4]]. 49 | - Hint 2: [About divn_eq] and [Search _ modn odd] 50 | #
# *) 51 | Lemma mod4Dsqr_even a b : (a ^ 2 + b ^ 2) %% 4 = 0 -> (~~ odd a) && (~~ odd b). 52 | Proof. 53 | (*D*)have sqr_x2Dy_mod4 x y : (x * 2 + y) ^ 2 = y ^ 2 %[mod 4]. 54 | (*D*) rewrite sqrnD addnAC mulnAC [2 * _]mulnC -mulnA -[2 * 2]/4. 55 | (*D*) by rewrite expnMn -[2 ^ 2]/4 -mulnDl -modnDml modnMl. 56 | (*D*)rewrite {1}(divn_eq a 2) {1}(divn_eq b 2) -modnDm. 57 | (*D*)by rewrite !sqr_x2Dy_mod4 modnDm !modn2; do 2!case: odd. 58 | (*A*)Qed. 59 | (** #
# 60 | - Deduce that if n is greater than 2 and a and b are solutions, then they are even. 61 | #
# *) 62 | Lemma sol_are_even n a b : n > 1 -> sol n a b -> (~~ odd a) && (~~ odd b). 63 | Proof. 64 | (*D*)case: n => [|[|n]]// _; rewrite /sol => /and3P[_ _ /eqP eq_a2Db]. 65 | (*D*)by rewrite mod4Dsqr_even// -eq_a2Db !expnS mulnA modnMr. 66 | (*A*)Qed. 67 | (** #
# 68 | - Prove that the solutions for n are the halves of the solutions for n + 2. 69 | 70 | - Hint: [Search _ odd double] and [Search _ "eq" "mul"]. 71 | 72 | #
# *) 73 | Lemma solSS n a b : sol n.+2 a b -> sol n a./2 b./2. 74 | Proof. 75 | (*D*)move=> soln2ab; have [//|a_even b_even] := andP (sol_are_even _ soln2ab). 76 | (*D*)rewrite /sol -[a]odd_double_half -[b]odd_double_half in soln2ab. 77 | (*D*)rewrite (negPf a_even) (negPf b_even) ?add0n ?double_gt0 in soln2ab. 78 | (*D*)rewrite /sol; move: soln2ab => /and3P[-> -> /=]. 79 | (*D*)by rewrite -(addn2 n) expnD -!muln2 !expnMn -mulnDl eqn_mul2r. 80 | (*A*)Qed. 81 | (** #
# 82 | - Prove there are no solutions for n even 83 | 84 | - Hint: Use [sol0] and [solSS]. 85 | #
# *) 86 | Lemma sol_even a b n : ~~ sol (2 * n) a b. 87 | Proof. 88 | (*D*)elim: n => [|n IHn] in a b *; first exact: sol0. 89 | (*D*)by apply/negP; rewrite mulnS => /solSS; apply/negP. 90 | (*A*)Qed. 91 | (** #
# 92 | - Certify the only solution when n is odd. 93 | 94 | - Hint 1: Use [sol1], [solSS] and [sol_are_even]. 95 | - Hint 2: Really sketch it on paper first! 96 | #
# *) 97 | Lemma sol_odd a b n : sol (2 * n + 1) a b = (a == 2 ^ n) && (b == 2 ^ n). 98 | Proof. 99 | (*D*)apply/idP/idP=> [|/andP[/eqP-> /eqP->]]; last first. 100 | (*D*) by rewrite /sol !expn_gt0/= expnD muln2 addnn -expnM mulnC. 101 | (*D*)elim: n => [|n IHn] in a b *; first by rewrite sol1. 102 | (*D*)rewrite mulnS !add2n !addSn => solab. 103 | (*D*)have [//|/negPf aNodd /negPf bNodd] := andP (sol_are_even _ solab). 104 | (*D*)rewrite /sol -[a]odd_double_half -[b]odd_double_half aNodd bNodd. 105 | (*D*)by rewrite -!muln2 !expnSr !eqn_mul2r IHn// solSS. 106 | (*A*)Qed. 107 | (** #
# 108 | *** Exercise 3: 109 | Certify the solutions of this problem. 110 | 111 | - Hint: Do not hesitate to take advantage of Coq's capabilities 112 | for brute force case analysis 113 | #
# *) 114 | Lemma ex3 n : (n + 4 %| 3 * n + 32) = (n \in [:: 0; 1; 6; 16]). 115 | Proof. 116 | (*D*)apply/idP/idP => [Hn|]; rewrite !inE; last first. 117 | (*D*) by move=> /or4P[] /eqP->. 118 | (*D*)have : n + 4 %| 3 * n + 32 - 3 * (n + 4) by rewrite dvdn_sub// dvdn_mull. 119 | (*D*)by rewrite mulnDr subnDl /= {Hn}; move: n; do 21?[case=>//]. 120 | (*A*)Qed. 121 | (** #
# 122 | *** Exercise 4: 123 | 124 | Certify the result of the euclidean division of 125 | #$$a b^n - 1\quad\textrm{ by }\quad b ^ {n+1}$$# 126 | 127 | #
# *) 128 | Lemma ex4 a b n : a > 0 -> b > 0 -> n > 0 -> 129 | edivn (a * b ^ n - 1) (b ^ n.+1) = 130 | ((a - 1) %/ b, ((a - 1) %% b) * b ^ n + b ^ n - 1). 131 | Proof. 132 | (*D*)move=> a_gt0 b_gt0 n_gt0; rewrite /divn modn_def. 133 | (*D*)have [q r aB1_eq r_lt] /= := edivnP (a - 1). 134 | (*D*)rewrite b_gt0 /= in r_lt. 135 | (*D*)have /(congr1 (muln^~ (b ^ n))) := aB1_eq. 136 | (*D*)rewrite mulnBl mulnDl mul1n. 137 | (*D*)move=> /(congr1 (addn^~ (b ^ n - 1))). 138 | (*D*)rewrite addnBA ?expn_gt0 ?b_gt0// subnK; last first. 139 | (*D*) by rewrite -[X in X <= _]mul1n leq_mul2r a_gt0 orbT. 140 | (*D*)rewrite -mulnA -expnS -addnA => ->. 141 | (*D*)rewrite edivn_eq addnBA ?expn_gt0 ?b_gt0//. 142 | (*D*)rewrite subnS subn0 prednK ?addn_gt0 ?expn_gt0 ?b_gt0 ?orbT//. 143 | (*D*)rewrite -[X in _ + X]mul1n -mulnDl addn1 expnS. 144 | (*D*)by rewrite leq_mul2r r_lt orbT. 145 | (*A*)Qed. 146 | (** #
# 147 | *** Exercise 5: 148 | 149 | Prove that the natural number interval #$$[n!+2\ ,\ n!+n]$$# 150 | contains no prime number. 151 | 152 | - Hint: Use [Search _ prime dvdn], [Search _ factorial], ... 153 | 154 | #
# *) 155 | Lemma ex5 n m : n`! + 2 <= m <= n`! + n -> ~~ prime m. 156 | Proof. 157 | (*D*)move=> m_in; move: (m_in); rewrite -[m](@subnKC n`!); last first. 158 | (*D*) by rewrite (@leq_trans (n`! + 2)) ?leq_addr//; by case/andP: m_in. 159 | (*D*)set k := (_ - _); rewrite !leq_add2l => /andP[k_gt1 k_le_n]. 160 | (*D*)apply/primePn; right; exists k. 161 | (*D*) by rewrite k_gt1/= -subn_gt0 addnK fact_gt0. 162 | (*D*)by rewrite dvdn_add// dvdn_fact// k_le_n (leq_trans _ k_gt1). 163 | (*A*)Qed. 164 | (** #
# *) -------------------------------------------------------------------------------- /SummerSchoolSophia/lesson6.v: -------------------------------------------------------------------------------- 1 | 2 | From mathcomp Require Import all_ssreflect. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | (** 9 | ---------------------------------------------------------- 10 | #
# 11 | ** Lesson 6: Summary 12 | 13 | - [have], [suff], [wlog] proof cuts, alternative to intermediate lemmas. 14 | - [set] and [pose] abstracting sub terms, 15 | - [-[pattern]/expression] substitutes an expression for a convertible one, 16 | - Good practices for real proofs. 17 | 18 | #
# 19 | 20 | ---------------------------------------------------------- 21 | #
# 22 | ** [have:] and [suff:] 23 | 24 | - [have i_items : statement.] asks you to prove [statement] first. 25 | - [suff i_items : statement.] asks you to prove [statement] first. 26 | - [have i_items := term.] genralizes [term] and puts in on the top of the stack. 27 | 28 | This last one is very useful as an alternative of [Check], since you can play with 29 | the result which is on the top of the stack. 30 | 31 | cf #ssreflect documentation on rewrite# 32 | 33 | #
# 34 | *) 35 | Lemma test_have n : 36 | reflect (exists x y z, x ^ n + y ^ n == z ^ n) ((n == 1) || (n == 2)). 37 | Proof. 38 | have test0 : forall x y z, x ^ 0 + y ^ 0 != z ^ 0 by []. 39 | have test1 : exists x y z, x ^ 1 + y ^ 1 = z ^ 1 by exists 0, 0, 0. 40 | have test2 : exists x y z, x ^ 2 + y ^ 2 = z ^ 2 by exists 3, 4, 5. 41 | suff test m : m >= 3 -> forall x y z, x ^ m + y ^ m != z ^ m. 42 | case: n => [|[|[|n]]]//=; constructor. 43 | - by move=> [x [y [z]]]; rewrite (negPf (test0 _ _ _)). 44 | - by have [x [y [z /eqP]]] := test1; exists x, y, z. 45 | - by have [x [y [z /eqP]]] := test2; exists x, y, z. 46 | - by move=> [x [y [z]]]; rewrite (negPf (test _ _ _ _ _)). 47 | (* The rest of the proof fits in a margin *) 48 | Abort. 49 | (** 50 | #
# 51 | 52 | 53 | #
# 54 | 55 | ---------------------------------------------------------- 56 | #
# 57 | ** Example: Infinitude of primes 58 | 59 | 60 | #
# 61 | *) 62 | Lemma prime_above (m : nat) : exists p, (prime p) && (m < p). 63 | Proof. 64 | have: 1 < m`! + 1 by rewrite addn1 ltnS fact_gt0. 65 | move=> /pdivP[q pr_q q_dv_m1]. 66 | exists q; rewrite pr_q /= ltnNge. 67 | apply: contraL q_dv_m1 => q_le_m. 68 | by rewrite dvdn_addr ?dvdn_fact ?prime_gt0 // gtnNdvd ?prime_gt1. 69 | Qed. 70 | 71 | (** 72 | #
# 73 | 74 | #


# 75 | 76 | #

# 77 | 78 | ---------------------------------------------------------- 79 | #
# 80 | ** Teaser: Order and max, a matter of symmetry 81 | 82 | #
# 83 | *) 84 | Lemma leq_max m n1 n2 : (m <= maxn n1 n2) = (m <= n1) || (m <= n2). 85 | Proof. 86 | case/orP: (leq_total n2 n1) => [le_n21|le_n12]. 87 | rewrite (maxn_idPl le_n21) orb_idr // => le_mn2. 88 | by apply: leq_trans le_mn2 le_n21. 89 | rewrite maxnC orbC. 90 | rewrite (maxn_idPl le_n12) orb_idr // => le_mn1. 91 | by apply: leq_trans le_mn1 le_n12. 92 | Abort. 93 | (** 94 | #
# 95 | 96 | Note the repetition, and the step dealing with symmetry. 97 | 98 | #
# 99 | 100 | ---------------------------------------------------------- 101 | #
# 102 | ** Without loss of generality 103 | 104 | - [wlog i_items : generalizations / H] in a goal G, lets you prove that 105 | - [(H -> G) -> G] as first goal 106 | - [H -> G] as a second goal. 107 | - its typical use is for breaking the symmetry. 108 | 109 | #
# 110 | *) 111 | Lemma leq_max m n1 n2 : (m <= maxn n1 n2) = (m <= n1) || (m <= n2). 112 | Proof. 113 | wlog le_n21: n1 n2 / n2 <= n1 => [th_sym|]. 114 | by case/orP: (leq_total n2 n1) => /th_sym; last rewrite maxnC orbC. 115 | rewrite (maxn_idPl le_n21) orb_idr // => le_mn2. 116 | by apply: leq_trans le_mn2 le_n21. 117 | Qed. 118 | (** 119 | #
# 120 | 121 | #


# 122 | 123 | #

# 124 | 125 | ---------------------------------------------------------- 126 | #
# 127 | ** [-[pattern]/term] lets you replace a term by a convertible one. 128 | 129 | e.g. 130 | - [-[2]/(1 + 1)] replaces [2] by [1 + 1], 131 | - [-[2 ^ 2]/4] replaces [2 ^ 2] by [4], 132 | - [-[m]/(0 * d + m)] replaces [m] by [0 * d + m]. 133 | 134 | Example: Euclidean division, simple and correct 135 | 136 | #
# 137 | *) 138 | Print edivn. 139 | Print edivn_rec. 140 | 141 | Lemma edivn_recE d m q : 142 | edivn_rec d m q = if m - d is m'.+1 then edivn_rec d m' q.+1 else (q,m). 143 | Proof. by case: m. Qed. 144 | 145 | Lemma edivnP m d (ed := edivn m d) : 146 | ((d > 0) ==> (ed.2 < d)) && (m == ed.1 * d + ed.2). 147 | Proof. 148 | case: d => [//=|d /=] in ed *. 149 | rewrite -[edivn m d.+1]/(edivn_rec d m 0) in ed *. 150 | rewrite -[m]/(0 * d.+1 + m). 151 | elim: m {-2}m 0 (leqnn m) @ed => [[]//=|n IHn [//=|m]] q le_mn. 152 | rewrite edivn_recE subn_if_gt; case: ifP => [le_dm|lt_md]; last first. 153 | by rewrite /= ltnS ltnNge lt_md eqxx. 154 | have /(IHn _ q.+1) : m - d <= n by rewrite (leq_trans (leq_subr d m)). 155 | by rewrite /= mulSnr -addnA -subSS subnKC. 156 | Qed. 157 | (** 158 | #
# 159 | 160 | #


# 161 | 162 | #

# 163 | 164 | ---------------------------------------------------------- 165 | #
# 166 | ** [set] and [pose] naming expressions 167 | 168 | - [pose name := pattern], generalizes every hole in the pattern, 169 | and puts a definition [name] in the context for it. 170 | It does NOT change the conclusion. 171 | - [set name := pattern], finds the pattern in the conclusion, 172 | and puts a definition [name] in the context. 173 | Finally, it substitutes the pattern for [name] in the conclution. 174 | 175 | #
# 176 | *) 177 | Lemma test m n k : (2 + n + 52) * m = k. 178 | Proof. 179 | pose x := (_ + n + _) * m. 180 | set y := (_ * _). 181 | Abort. 182 | (** 183 | #
# 184 | 185 | #
# 186 | 187 | ---------------------------------------------------------- 188 | #
# 189 | ** Good practices for real proofs. 190 | 191 | - Never unfold a definition from a library, unless you really mean it. 192 | (Unfolding will happen mostly for the symbols YOU defined) 193 | - Take piece of paper and sketch the proof with a pen, 194 | do NOT let the proof assitant 195 | guide you, only let it assist you. 196 | - Take a look at the header of the library you want to use, to know the concepts. 197 | - Use [Search _ symbol1 symbol2] to lookup the lemmas about these concepts. 198 | - Read the name conventions and use them in your own proofs and searches. 199 | 200 | #
# 201 | 202 | ---------------------------------------------------------- 203 | #
# 204 | ** A little tour of the [div] library 205 | 206 | If time allows: 207 | - [dvdn], [modn], [divn], [m = n %[mod d]] ... 208 | - juggling with specificies of division by 2. 209 | #
# 210 | 211 | *) 212 | -------------------------------------------------------------------------------- /SummerSchoolSophia/lesson8.v: -------------------------------------------------------------------------------- 1 | From HB Require Import structures. 2 | From mathcomp Require Import all_ssreflect. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | (** 9 | 10 | ---------------------------------------------------------- 11 | #
# 12 | ** Lesson 8: summary 13 | 14 | - OOP with mix-ins 15 | - subtypes & automation 16 | 17 | Let's remember the truth: 18 | 19 | #
# 20 | Coq is an object oriented 21 | programming language. 22 | #
# 23 | 24 | 25 | #
# 26 | ---------------------------------------------------------- 27 | #
# 28 | ** OOP 29 | 30 | # # 31 | 32 | Let's see another interface, the one of finite types 33 | 34 | #
# 35 | *) 36 | 37 | HB.about Finite. 38 | HB.about isFinite. 39 | 40 | Print finite_axiom. 41 | Print count_mem. 42 | 43 | Eval lazy in count_mem 3 [:: 1;2;3;4;3;2;1]. 44 | 45 | (* The property of finite types is that *) 46 | 47 | Check fun (T : eqType) (enum : seq T) => 48 | forall x : T, count_mem x enum = 1. 49 | 50 | Section Example. 51 | 52 | Variable T : finType. 53 | 54 | (* Cardinality of a finite type *) 55 | Check #| T |. 56 | 57 | (* "bounded" quantification *) 58 | Check [forall x : T, x == x] && false. 59 | Fail Check (forall x : T, x == x) && false. 60 | 61 | End Example. 62 | 63 | 64 | (** 65 | #
# 66 | 67 | 68 | #
(notes)
# 69 | This slide corresponds to 70 | section 6.1 and 6.2 of 71 | #the Mathematical Components book# 72 | #
# 73 | 74 | #


# 75 | #

# 76 | 77 | 78 | 79 | 80 | ---------------------------------------------------------- 81 | #
# 82 | ** Sub types 83 | 84 | A sub type extends another type by adding a property. 85 | The new type has a richer theory. 86 | The new type inherits the original theory. 87 | 88 | Let's define the type of homogeneous tuples 89 | 90 | #
# 91 | *) 92 | 93 | Module Tup. 94 | 95 | Record tuple_of (n : nat) T := Tuple { 96 | tval :> seq T; 97 | tsize : size tval == n 98 | }. 99 | Notation "n .-tuple" := (tuple_of n) : type_scope. 100 | 101 | Lemma size_tuple T n (t : n .-tuple T) : size t = n. 102 | Proof. by case: t => s /= /eqP. Qed. 103 | 104 | Example seq_on_tuple (n : nat) (t : n .-tuple nat) : 105 | size (rev [seq 2 * x | x <- rev t]) = size t. 106 | Proof. 107 | by rewrite map_rev revK size_map. 108 | Undo. 109 | rewrite size_tuple. 110 | Fail rewrite size_tuple. 111 | Abort. 112 | 113 | 114 | (** 115 | #
# 116 | 117 | We instrument Coq to automatically promote 118 | sequences to tuples. 119 | 120 | #
# 121 | *) 122 | 123 | Lemma rev_tupleP n A (t : n .-tuple A) : size (rev t) == n. 124 | Proof. by rewrite size_rev size_tuple. Qed. 125 | Canonical rev_tuple n A (t : n .-tuple A) := Tuple (rev_tupleP t). 126 | 127 | Lemma map_tupleP n A B (f: A -> B) (t: n .-tuple A) : size (map f t) == n. 128 | Proof. by rewrite size_map size_tuple. Qed. 129 | Canonical map_tuple n A B (f: A -> B) (t: n .-tuple A) := Tuple (map_tupleP f t). 130 | 131 | Example seq_on_tuple2 n (t : n .-tuple nat) : 132 | size (rev [seq 2 * x | x <- rev t]) = size t. 133 | Proof. rewrite size_tuple. rewrite size_tuple. by []. Qed. 134 | 135 | (** 136 | #
# 137 | 138 | Reamrk how [t] is a tuple, then it becomes a list by going 139 | trough rev and map, and is finally "promoted" back to a tuple 140 | by this [Canonical] magic. 141 | 142 | 143 | Now we the tuple type to form an eqType, 144 | exactly as seq does. 145 | 146 | Which is the expected comparison for tuples? 147 | 148 | #
# 149 | *) 150 | 151 | Lemma p1 : size [:: 1;2] == 2. Proof. by []. Qed. 152 | Lemma p2 : size ([:: 1] ++ [::2]) == 2. Proof. by rewrite cat_cons cat0s. Qed. 153 | 154 | Definition t1 := {| tval := [::1;2]; tsize := p1 |}. 155 | Definition t2 := {| tval := [::1] ++ [::2]; tsize := p2 |}. 156 | 157 | Lemma tuple_uip : t1 = t2. 158 | Proof. 159 | rewrite /t1 /t2. rewrite /=. 160 | Fail by []. 161 | congr (Tuple _). 162 | Fail by []. 163 | (*About bool_irrelevance.*) 164 | apply: bool_irrelevance. 165 | Qed. 166 | 167 | (** 168 | #
# 169 | 170 | Given that propositions are expressed (whenever possible) 171 | as booleans we can systematically prove that proofs 172 | of these properties are irrelevant. 173 | 174 | As a consequence we can form subtypes and systematically 175 | prove that the projection to the supertype is injective, 176 | that means we can craft an eqType. 177 | 178 | #
# 179 | *) 180 | 181 | HB.instance Definition _ (n : nat) T := [isSub for (@tval n T)]. 182 | HB.instance Definition _ n (T : eqType) := [Equality of n .-tuple T by <:]. 183 | 184 | Check [eqType of 3.-tuple nat]. 185 | 186 | Example test_eqtype (x y : 3.-tuple nat) : x == y -> True. 187 | Proof. 188 | move=> /eqP H. 189 | Abort. 190 | 191 | End Tup. 192 | 193 | (** 194 | #
# 195 | 196 | Tuples is are part of the library, that also contains 197 | many other "promotions" 198 | 199 | #
# 200 | *) 201 | 202 | Check [finType of 3.-tuple bool]. 203 | Fail Check [finType of 3.-tuple nat]. 204 | 205 | (** 206 | #
# 207 | 208 | Tuples is not the only subtype part of the library. 209 | Another one is ['I_n], the finite type of natural 210 | numbers smaller than n. 211 | 212 | #
# 213 | *) 214 | 215 | Print ordinal. 216 | Check [eqType of 'I_3]. 217 | Check [finType of 'I_3]. 218 | 219 | About tnth. (* like the safe nth function for vectors *) 220 | 221 | (** 222 | #
# 223 | 224 | It is easy to combine these bricks by subtyping (and "specialization") 225 | 226 | #
# 227 | *) 228 | 229 | Check {set 'I_4} : Type. 230 | Check forall a : {set 'I_4}, (a == set0) || (1 < #| a | < 4). 231 | Print set_type. 232 | Check {ffun 'I_4 -> bool} : Type. 233 | Check [eqType of #| 'I_4 | .-tuple bool]. 234 | Check [finType of #| 'I_4 | .-tuple bool]. 235 | 236 | Check {ffun 'I_4 * 'I_6 -> nat} : Type. 237 | Check [eqType of {ffun 'I_4 * 'I_6 -> nat}] : Type. 238 | 239 | From mathcomp Require Import all_algebra. 240 | Open Scope ring_scope. 241 | 242 | Print matrix. 243 | 244 | Section Rings. 245 | 246 | Variable R : ringType. 247 | 248 | Check forall x : R, x * 1 == x. 249 | 250 | Check forall m : 'M[R]_(4,4), m == m * m. 251 | 252 | End Rings. 253 | 254 | (** 255 | #
# 256 | 257 | #
(notes)
# 258 | This slide corresponds to 259 | section 6.1 and 6.2 of 260 | #the Mathematical Components book# 261 | #
# 262 | 263 | #


# 264 | #

# 265 | 266 | ---------------------------------------------------------- 267 | #
# 268 | ** Sum up 269 | 270 | - subtypes add properties and inherit the theory of the supertype 271 | thanks to boolean predicates (UIP). 272 | In some cases the property can be inferred by Coq, letting one apply 273 | a lemma about the subtype on terms of the supertype. 274 | 275 | 276 | #
# 277 | 278 | *) 279 | -------------------------------------------------------------------------------- /SummerSchoolSophia/lesson5.v: -------------------------------------------------------------------------------- 1 | 2 | From mathcomp Require Import all_ssreflect. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | (** 9 | 10 | ---------------------------------------------------------- 11 | #
# 12 | ** Lesson 5 : summary 13 | 14 | - bool vs Prop 15 | - views as iff 16 | - using views 17 | - reflect and other inductive "spec" 18 | 19 | #
# 20 | 21 | 22 | ---------------------------------------------------------- 23 | #
# 24 | ** Propositions and booleans 25 | 26 | So far we used boolean connectives. 27 | But this is not what is used in the Curry Howard 28 | correspondence to represent connectives and their 29 | proofs. 30 | 31 | #
# 32 | *) 33 | Print andb. 34 | Print and. 35 | 36 | Print orb. 37 | Print or. 38 | 39 | Check true ==> false. 40 | Check True -> False. 41 | 42 | (** 43 | #
# 44 | 45 | Let's play a little with [and] and [andb], [or] and [orb] 46 | in order to understand the difference. 47 | 48 | #
# 49 | *) 50 | 51 | 52 | Lemma test_and (P : bool) : 53 | True /\ P -> P. (* true && P -> P. *) 54 | Proof. 55 | move=> t_p. 56 | case: t_p => t p. 57 | apply: p. 58 | Qed. 59 | 60 | Lemma test_orb (P Q R : bool) : 61 | P \/ Q -> R. (* P || Q -> R *) 62 | Proof. 63 | move=> p_q. 64 | case: p_q. 65 | Abort. 66 | 67 | 68 | (** 69 | #
# 70 | 71 | Propositions: 72 | - structure to your proof as a tree 73 | - more expressive logic (closed under forall, exists...) 74 | 75 | Booleans: 76 | - computation & Excluded Middle 77 | - Uniqueness of Identity Proofs (lesson 4) 78 | 79 | We want the best of the two worlds, and a way to 80 | link them: views. 81 | 82 | #


# 83 | 84 | #

(notes)
# 85 | 86 | This slide corresponds to 87 | section 3.x of 88 | #the Mathematical Components book# 89 | #
# 90 | 91 | 92 | #
# 93 | 94 | ---------------------------------------------------------- 95 | #
# 96 | ** Stating and proving a reflection view 97 | 98 | To link a concept in bool and one in Prop we typically 99 | use the [reflect] predicate. 100 | 101 | To prove [reflect] we use the [iffP] lemma that 102 | turns it into a double implication. This is also a 103 | recap of many of the proof commands we have seen 104 | so far. 105 | 106 | #
# 107 | *) 108 | 109 | About iffP. 110 | About idP. 111 | 112 | Lemma eqnP {n m : nat} : 113 | reflect (n = m) (eqn n m). 114 | Proof. 115 | apply: (iffP idP). 116 | elim: n m => [|n IH] m; case: m => // m Hm. 117 | by rewrite (IH m Hm). 118 | move=> def_n; rewrite {}def_n. 119 | Undo. 120 | move=> ->. (* move + rewrie + clear idiom *) 121 | by elim: m. 122 | Qed. 123 | 124 | (** 125 | #
# 126 | 127 | #
(notes)
# 128 | This slide corresponds to 129 | sections 4.1.1 and 4.1.2 of 130 | #the Mathematical Components book# 131 | #
# 132 | 133 | 134 | #
# 135 | 136 | 137 | ---------------------------------------------------------- 138 | #
# 139 | ** Using views 140 | 141 | The syntax [/view] can be put in intro patterns 142 | to modify the top assumption using [view] 143 | 144 | #
# 145 | *) 146 | About andP. 147 | 148 | Lemma example n m k : k <= n -> 149 | (n <= m) && (m <= k) -> n = k. 150 | Proof. 151 | move=> lekn /andP H; case: H => lenm lemk. 152 | Undo. 153 | move=> lekn /andP[lenm lemk]. (* view + case idiom *) 154 | Abort. 155 | 156 | (** 157 | #
# 158 | 159 | The [apply:] tactic accepts a [/view] flag 160 | to modify the goal using [view]. 161 | 162 | #
# 163 | *) 164 | 165 | Lemma leq_total m n : (m <= n) || (m >= n). 166 | Proof. 167 | (* About implyNb.*) 168 | rewrite -implyNb -ltnNge. 169 | apply/implyP. 170 | (* About ltnW *) 171 | by apply: ltnW. 172 | Qed. 173 | 174 | (** 175 | #
# 176 | 177 | The [case:] tactic accepts a [/view] flag 178 | to modify the term being analyzed just before 179 | performing the case analysis. 180 | 181 | #
# 182 | *) 183 | 184 | Lemma leq_max m n1 n2 : 185 | (m <= maxn n1 n2) = (m <= n1) || (m <= n2). 186 | Proof. 187 | case/orP: (leq_total n2 n1) => [le_n21 | le_n12]. 188 | Abort. 189 | 190 | (** 191 | #
# 192 | 193 | #


# 194 | 195 | #

(notes)
# 196 | This slide corresponds to 197 | sections 4.1.3 and 4.1.4 of 198 | #the Mathematical Components book# 199 | #
# 200 | 201 | 202 | #
# 203 | 204 | 205 | ---------------------------------------------------------- 206 | #
# 207 | ** The reflect predicate and other specs 208 | 209 | The [reflect] inductive predicate has an index. 210 | 211 | Indexes are replaced by the value dictated by the 212 | constructor when performing a case analysis. In a way 213 | indexes express equations that are substituted 214 | automatically. 215 | 216 | In Mathematical Components we exploit this feature 217 | of the logic in order to structure proofs that proceed 218 | by case split. 219 | 220 | 221 | #
# 222 | *) 223 | Print Bool.reflect. 224 | About andP. 225 | 226 | Lemma example_spec a b : a && b ==> (a == b). 227 | Proof. 228 | by case: andP => [ [-> ->] | // ]. 229 | Qed. 230 | 231 | (** 232 | #
# 233 | 234 | Note that [(andP _ _)] has in the type, as the value of 235 | the index, [(_ && _)] that we call "a pattern". The [case:] 236 | tactic looks for a subterm of the goal that matches the 237 | pattern and "guesses" that the two [_] are respectively 238 | [a] and [b]. This form of automation is the same of [rewrite]. 239 | 240 | One can craft many "spec" to model the structure of a 241 | proof. Eg [leqP] and [ltngtP] 242 | 243 | #
# 244 | *) 245 | 246 | About leqP. 247 | Print leq_xor_gtn. 248 | 249 | Lemma example_spec2 a b : (a <= b) || (b < a). 250 | Proof. 251 | case: leqP. 252 | Abort. 253 | 254 | About ltngtP. 255 | Print compare_nat. 256 | 257 | Lemma example_spec3 a b : (a <= b) || (b < a) || (b == a). 258 | Proof. 259 | case: ltngtP. 260 | Abort. 261 | 262 | 263 | (** 264 | #
# 265 | 266 | 267 | ---------------------------------------------------------- 268 | 269 | #
# 270 | ** Using views with non reflexive lemmas 271 | 272 | - By processing an assumption through a lemma. 273 | - The leading / makes the lemma work as a function. 274 | - If lemma states A -> B, we ca use it as a function to get a proof of B from a proof of A. 275 | - One can also chain multiple views on the same stack item. 276 | 277 | 278 | #
# 279 | *) 280 | 281 | About prime_gt1. 282 | 283 | Lemma example_2 x y : prime x -> odd y -> 2 < y + x. 284 | Proof. 285 | move/prime_gt1 => x_gt_1. (* view through prime_gt1 *) 286 | Undo. 287 | move/prime_gt1/ltnW. 288 | 289 | Abort. 290 | 291 | 292 | (** 293 | #
# 294 | Using views with tactics 295 | - we can also use views on non reflexive lemmas: apply/V, case/V. 296 | 297 | #
# 298 | *) 299 | 300 | 301 | 302 | (** 303 | #
# 304 | The view mechanism can also be used for double implication: 305 | - A <-> B is a notation for the conjonction: (A -> B) /\ (B -> A) 306 | - If V is a proof of (A <-> B), apply/V will automatically guess which part of the conjonction it will use 307 | 308 | 309 | 310 | #
# 311 | *) 312 | 313 | Lemma example_3 A B (V: A <-> B): B -> A . 314 | Proof. 315 | move=>b. 316 | apply/V. 317 | by []. 318 | Qed. 319 | 320 | 321 | 322 | 323 | 324 | 325 | (** 326 | #
# 327 | 328 | #


# 329 | 330 | 331 | #

(notes)
# 332 | This slide corresponds to 333 | section 4.2 of 334 | #the Mathematical Components book# 335 | #
# 336 | 337 | 338 | #
# 339 | 340 | ---------------------------------------------------------- 341 | #
# 342 | ** Lesson 5: sum up 343 | 344 | - [reflect] and [iffP] 345 | - [case/v: m] case split + view application 346 | - [move=> /v H] introduction + view 347 | - [apply/v: t] backchain with a view 348 | - [_spec] predicates to model case split [leqP] and [ifP] 349 | - [move=> []] (for case) 350 | - [move=> ->] (for rewrite) 351 | - [... => ...] the arrow can be used after any command 352 | 353 | #
# 354 | 355 | *) 356 | -------------------------------------------------------------------------------- /AnIntroductionToSmallScaleReflectionInCoq/section5.v: -------------------------------------------------------------------------------- 1 | (******************************************************************************) 2 | (* Solutions of exercises : Type inference by canonical structures *) 3 | (******************************************************************************) 4 | 5 | 6 | From HB Require Import structures. 7 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq path. 8 | From mathcomp Require Import choice fintype tuple finset. 9 | 10 | Set Implicit Arguments. 11 | Unset Strict Implicit. 12 | Import Prenex Implicits. 13 | 14 | (******************************************************************************) 15 | (* Exercise 5.1.1 *) 16 | (******************************************************************************) 17 | (* 2023 This has been updated to use HB *) 18 | 19 | HB.mixin Record IsZmodule T := { 20 | zero : T; 21 | opp : T -> T; 22 | add : T -> T -> T; 23 | addA : associative add; 24 | addC: commutative add; 25 | addm0 : left_id zero add; 26 | add0m : left_inverse zero opp add 27 | }. 28 | 29 | #[short(type="zmodule")] 30 | HB.structure Definition Zmodule := { T of IsZmodule T }. 31 | 32 | #[verbose] 33 | HB.instance Definition Bool_Zmodule := 34 | IsZmodule.Build bool addbA addbC addFb addbb. 35 | 36 | Notation "x \+ y" := (add x y)(at level 50,left associativity). 37 | 38 | (* We first need to prove that zmadd is associative and commutative *) 39 | (* The proof consists in breaking successivly the two nested records *) 40 | (*to recover all the ingredients present in the zmodule_mixin. Then *) 41 | (*the goal becomes trivial because the associative and commutative *) 42 | (*requirements were present in the spec. *) 43 | (* Update with HB, associativity and commutativity are available directly *) 44 | 45 | Lemma zmaddA : forall m : zmodule, associative (@add m). 46 | Proof. move=> m; apply: addA. Qed. 47 | 48 | Lemma zmaddC : forall m : zmodule, commutative (@add m). 49 | Proof. move=> m; apply: addC. Qed. 50 | 51 | 52 | (* No we can conveniently prove the lemma *) 53 | (* The ssreflect rewrite tactic allows rewrite redex selection by *) 54 | (* pattern, and this is used here to select the occurrence where *) 55 | (* commutativity should be used.*) 56 | Lemma zmaddAC : forall (m : zmodule)(x y z : m), x \+ y \+ z = x \+ z \+ y. 57 | Proof. 58 | by move=> M x y z; rewrite -zmaddA [y \+ _]zmaddC zmaddA. 59 | Qed. 60 | 61 | (******************************************************************************) 62 | (* Exercise 5.2.1 *) 63 | (******************************************************************************) 64 | 65 | (* Be aware that the Print command shows terms as they are represented *) 66 | (*by the system, which is possibliy syntactically slightly different *) 67 | (*from the definition typed by the user (specially in the case of *) 68 | (*nested pattern matching *) 69 | 70 | HB.about hasDecEq. 71 | HB.about hasDecEq.eq_op. 72 | HB.about hasDecEq.eqP. 73 | 74 | Print eqn. 75 | Check @eqnP. 76 | 77 | Print eqb. 78 | Check @eqbP. 79 | 80 | (* Look for nat, bool, Equality.sort in the answer of the command: *) 81 | Print Canonical Projections. 82 | (* The equations stored after these declarations are respectively : 83 | eqn <- hasDecEq.eq_op ( ssrnat.HB_unnamed_factory_1 ) 84 | BinNat.N.eqb <- hasDecEq.eq_op ( ssrnat.HB_unnamed_factory_3 ) 85 | eqb <- hasDecEq.eq_op ( HB_unnamed_factory_4 ) 86 | 87 | [ Equality.sort ? == nat ] => ? = nat_eqType 88 | [ Equality.sort ? == bool ] => ? = bool_eqType 89 | *) 90 | 91 | 92 | (******************************************************************************) 93 | (* Exercise 5.2.2 *) 94 | (******************************************************************************) 95 | 96 | (* This script uses a complex intro pattern: the apply: (iffP andP) *) 97 | (*tactic breaks the reflect goal into two subgoals for each *) 98 | (*implication, while interpreting the first assumption of each *) 99 | (*generated sugoal with the andP view. Then the [[] | [<- <- ]] // *) 100 | (*intropattern simultaneously describes the introduction operations *) 101 | (*performed on the two subgoals:*) 102 | (*- on the first subgoal the [] casing intropattern splits the *) 103 | (*conjunction hypothesis into two distinct assumptions.*) 104 | (* - on the second hypothesis the [<- <- ] is composed of a [] *) 105 | (*casing intropattern and two <- rewrite intropatterns. [] performs an *) 106 | (*injection on the equality, generating two equalities, one for each *) 107 | (*components of the pairs. Then both equalities are rewritten form *) 108 | (*left to right by <- <-. This subgoal is now trivial hence closed by *) 109 | (*// *) 110 | Lemma tuto_pair_eqP : forall T1 T2, Equality.axiom (@pair_eq T1 T2). 111 | Proof. 112 | move=> T1 T2 [u1 u2] [v1 v2] /=. 113 | apply: (iffP andP) => [[]|[<- <-]] //. 114 | by do 2!move/eqP->. 115 | Qed. 116 | 117 | 118 | (******************************************************************************) 119 | (* Exercise 5.3.1 *) 120 | (******************************************************************************) 121 | Section SeqMem. 122 | 123 | Variable T : eqType. 124 | 125 | Implicit Type s : seq T. 126 | Implicit Types x y : T. 127 | 128 | Lemma tuto_in_cons : forall y s x, 129 | (x \in y :: s) = (x == y) || (x \in s). 130 | Proof. by []. Qed. 131 | 132 | (* by [] is an alternative syntax for the done tactic *) 133 | Lemma tuto_in_nil : forall x, (x \in [::]) = false. 134 | Proof. by []. Qed. 135 | 136 | Lemma tuto_mem_seq1 : forall x y, (x \in [:: y]) = (x == y). 137 | Proof. by move=> x y; rewrite in_cons orbF. Qed. 138 | 139 | (* Here we do not even need to name the elements introduced since they *) 140 | (*will never be used later in the script. We let the system choose *) 141 | (*names using the move=> * tactic. *) 142 | Lemma tuto_mem_head : forall x s, x \in x :: s. 143 | Proof. by move=> *; exact: predU1l. Qed. 144 | 145 | (******************************************************************************) 146 | (* Exercise 5.3.2 *) 147 | (******************************************************************************) 148 | 149 | Lemma tuto_mem_cat : forall x s1 s2, 150 | (x \in s1 ++ s2) = (x \in s1) || (x \in s2). 151 | Proof. 152 | by move=> x s1 s2; elim: s1 => //= y s1 IHs; rewrite !inE /= -orbA -IHs. 153 | Qed. 154 | 155 | Lemma tuto_mem_behead: forall s, {subset behead s <= s}. 156 | Proof. move=> [|y s] x //; exact: predU1r. Qed. 157 | 158 | Fixpoint tuto_has (a : pred T) s := 159 | if s is x :: s' then a x || tuto_has a s' else false. 160 | 161 | Lemma tuto_hasP : forall (a : pred T) s, 162 | reflect (exists2 x, x \in s & a x) (has a s). 163 | Proof. 164 | move=> a; elim=> [|y s IHs] /=; first by right; case. 165 | case ay: (a y); first by left; exists y; rewrite ?mem_head. 166 | apply: (iffP IHs) => [] [x ysx ax]; exists x => //; first exact: mem_behead. 167 | Search _ (_ \in cons _ _). 168 | by move: ysx ax; rewrite in_cons; case/orP=> //; move/eqP->; rewrite ay. 169 | Qed. 170 | 171 | (* In fact, the last line of the previous script can be simplified by *) 172 | (*the use of the predU1P view lemma. This is possible since the mem *) 173 | (*function on sequences is exactly programmed as required by the statement *) 174 | (*of predU1P. *) 175 | 176 | 177 | Lemma tuto_hasP_alt: forall (a : pred T) s, 178 | reflect (exists2 x, x \in s & a x) (has a s). 179 | Proof. 180 | move=> a; elim=> [|y s IHs] /=; first by right; case. 181 | case ay: (a y); first by left; exists y; rewrite ?mem_head. 182 | apply: (iffP IHs) => [] [x ysx ax]; exists x => //; first exact: mem_behead. 183 | by case: (predU1P ysx) ax => [->|//]; rewrite ay. 184 | Qed. 185 | 186 | 187 | Fixpoint tuto_all a s := if s is x :: s' then a x && tuto_all a s' else true. 188 | 189 | (* We again use predU1P to shortcut the combination of incons and eqP *) 190 | Lemma tuto_allP : forall (a : pred T) s, 191 | reflect (forall x, x \in s -> a x) (all a s). 192 | Proof. 193 | Proof. 194 | move=> a; elim=> [|x s IHs]; first by left. 195 | rewrite /= andbC; case: IHs => IHs /=. 196 | apply: (iffP idP) => [Hx y|]; last by apply; exact: mem_head. 197 | by case/predU1P=> [->|Hy]; auto. 198 | by right; move=> H; case IHs; move=> y Hy; apply H; exact: mem_behead. 199 | Qed. 200 | 201 | End SeqMem. -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise5_solution.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | 3 | Module easy. 4 | 5 | Implicit Type p q r : bool. 6 | Implicit Type m n a b c : nat. 7 | 8 | (** *** Exercise 1: 9 | - look for lemmas supporting contrapositive reasoning 10 | - use the eqP view to finish the proof. 11 | *) 12 | Lemma bool_gimmics1 a : a != a.-1 -> a != 0. 13 | (*D*)Proof. by apply: contra; move => /eqP E; rewrite E. Qed. 14 | 15 | (** *** Exercise 2: 16 | - it helps to find out what is behind [./2] and [.*2] in order to [Search] 17 | - any proof would do, but there is one not using [implyP] 18 | *) 19 | Lemma view_gimmics1 p a b : p -> (p ==> (a == b.*2)) -> a./2 = b. 20 | (*D*)Proof. by move=> -> /eqP->; exact: doubleK. Qed. 21 | 22 | (** *** Exercise 3: 23 | - Prove this view by unfolding maxn and then using [leqP] 24 | *) 25 | Lemma maxn_idPl m n : reflect (maxn m n = m) (m >= n). 26 | Proof. apply: (iffP idP). 27 | (*D*)by rewrite /maxn; case: leqP. 28 | (*D*)by move=> <-; rewrite /maxn; case: leqP. 29 | Qed. 30 | 31 | (** *** Exercise 4: 32 | - there is no need to prove [reflect] with [iffP]: here just use [rewrite] and [apply] 33 | - check out the definitions and theory of [leq] and [maxn] 34 | - proof sketch: 35 | << 36 | n <= m = n - m == 0 37 | = m + n - m == m + 0 38 | = maxn m n == m 39 | >> *) 40 | Lemma maxn_idPl_bis m n : reflect (maxn m n = m) (m >= n). 41 | (*D*)Proof. by rewrite -subn_eq0 -(eqn_add2l m) addn0 -maxnE; apply: eqP. Qed. 42 | 43 | End easy. 44 | 45 | 46 | 47 | Module reflect1. 48 | 49 | 50 | (** *** Exercise 5: 51 | - Prove some reflection lemmas for the proposed reflect definition 52 | *) 53 | 54 | (** 55 | - A possible definition for reflect 56 | *) 57 | 58 | Inductive reflectl (P : Prop) b := 59 | | RT (p : P) (e : b = true) 60 | | RF (p : ~ P) (e : b = false). 61 | About reflect. 62 | 63 | 64 | Lemma andP (b1 b2 : bool) : reflectl (b1 /\ b2) (b1 && b2). 65 | Proof. 66 | (*D*)by case: b1; case: b2; [ left | right => //= [[l r]] ..]. 67 | Qed. 68 | 69 | Lemma orP (b1 b2 : bool) : reflectl (b1 \/ b2) (b1 || b2). 70 | Proof. 71 | (*D*)case: b1; case: b2; [ left; by [ move | left | right ] .. |]. 72 | (*D*)by right=> // [[l|r]]. 73 | Qed. 74 | 75 | Lemma implyP (b1 b2 : bool) : reflectl (b1 -> b2) (b1 ==> b2). 76 | Proof. 77 | (*D*)by case: b1; case: b2; [ left | right | left ..] => //= /(_ isT). 78 | Qed. 79 | 80 | End reflect1. 81 | 82 | 83 | 84 | 85 | 86 | (** *** Exercise 6: 87 | - Get excluded-middle when P is equivalent to a "bool" "decidable" 88 | *) 89 | 90 | (** Equivalence definition *) 91 | 92 | Definition bool_Prop_equiv (P : Prop) (b : bool) := b = true <-> P. 93 | Lemma test_bool_Prop_equiv b P : bool_Prop_equiv P b -> P \/ ~ P. 94 | Proof. 95 | (*D*)case: b; case => hlr hrl. 96 | (*D*) by left; apply: hlr. 97 | (*D*)by right => hP; move: (hrl hP). 98 | Qed. 99 | 100 | (** *** Exercise 7: 101 | - Let's use standard reflect (and iffP, idP etc...) 102 | *) 103 | Lemma iffP_lr (P : Prop) (b : bool) : 104 | (P -> b) -> (b -> P) -> reflect P b. 105 | Proof. 106 | (*D*)by move=> *; apply: (iffP idP). 107 | Qed. 108 | 109 | Lemma iffP_rl (P : Prop) (b : bool) : 110 | reflect P b -> ((P -> b) /\ (b -> P)). 111 | Proof. 112 | (*D*) by case: b; case=> p; split. 113 | Qed. 114 | 115 | 116 | Lemma eqnP {n m : nat} : 117 | reflect (n = m) (eqn n m). 118 | Proof. 119 | (*D*)apply: (iffP idP) => [|->]; last by elim: m. 120 | (*D*)by elim: n m => [[]|n IH [//|m] /IH ->]. 121 | Qed. 122 | 123 | 124 | (** *** Exercise 8: 125 | If you have time.. more reflections 126 | 127 | - leq_max : use leq_total maxn_idPl 128 | - dvdn_fact: use leq_eqVlt dvdn_mulr dvdn_mull 129 | *) 130 | 131 | Lemma nat_inj_eq T (f : T -> nat) x y : 132 | injective f -> 133 | reflect (x = y) (eqn (f x) (f y)). 134 | Proof. 135 | (*D*)by move=> f_inj; apply: (iffP eqnP) => [/f_inj|-> //]. 136 | Qed. 137 | 138 | Lemma leq_max m n1 n2 : (m <= maxn n1 n2) = (m <= n1) || (m <= n2). 139 | Proof. 140 | (*D*)wlog le_n21: n1 n2 / n2 <= n1 => [th_sym|]. 141 | (*D*) by case/orP: (leq_total n2 n1) => /th_sym; last rewrite maxnC orbC. 142 | (*D*)by rewrite (maxn_idPl le_n21) orb_idr // => /leq_trans->. 143 | Qed. 144 | 145 | 146 | Lemma dvdn_fact m n : 0 < m <= n -> m %| n`!. 147 | Proof. 148 | (*D*)case: m => //= m; elim: n => //= n IHn; rewrite ltnS leq_eqVlt. 149 | (*D*)by case/orP=> [/eqP-> | /IHn]; [apply: dvdn_mulr | apply: dvdn_mull]. 150 | Qed. 151 | 152 | Lemma prime_above m : exists2 p, m < p & prime p. 153 | Proof. 154 | (*D*)have /pdivP[p pr_p p_dv_m1]: 1 < m`! + 1 by rewrite addn1 ltnS fact_gt0. 155 | (*D*)exists p => //; rewrite ltnNge; apply: contraL p_dv_m1 => p_le_m. 156 | (*D*)by rewrite dvdn_addr ?dvdn_fact ?prime_gt0 // gtnNdvd ?prime_gt1. 157 | Qed. 158 | 159 | (** 160 | For this section: 161 | only move=> h, move/V: h => h, case/V: h, by ... allowed 162 | *) 163 | 164 | Goal forall (P Q : Prop), (P <-> Q) -> P -> Q. 165 | Proof. 166 | (*D*) move=> P Q h hp. by move/h: hp => hp. 167 | Qed. 168 | 169 | Goal forall (P : nat -> Prop) (Q : Prop), 170 | (P 0 -> Q) 171 | -> (forall n, P n.+1 -> P n) 172 | -> P 4 -> Q. 173 | Proof. 174 | (*D*) move=> P Q P0Q Hp. 175 | (*D*) by move/Hp/Hp/Hp/Hp. 176 | Qed. 177 | 178 | 179 | (** No case analysis on b, b1, b2 allowed *) 180 | Goal forall (b b1 b2 : bool), (b1 -> b) -> (b2 -> b) -> b1 || b2 -> b. 181 | Proof. 182 | (*D*) move=> b b1 b2 h1 h2 h. case/orP: h. 183 | (*D*) by move/h1. 184 | (*D*) by move/h2. 185 | Qed. 186 | 187 | Goal forall (Q : nat -> Prop) (p1 p2 : nat -> bool) x, 188 | ((forall y, Q y -> p1 y /\ p2 y) /\ Q x) -> p1 x && p2 x. 189 | Proof. 190 | (*D*) move=> Q p1 p2 x h. case: h. move=> /(_ x) h qx. 191 | (*D*) move/h: qx. by move/andP. 192 | Qed. 193 | 194 | Goal forall (Q : nat -> Prop) (p1 p2 : nat -> bool) x, 195 | ((forall y, Q y -> p1 y \/ p2 y) /\ Q x) -> p1 x || p2 x. 196 | Proof. 197 | (*D*) move=> Q p1 p2 x h. case: h. move=> /(_ x) h qx. 198 | (*D*) move/h: qx. by move/orP. 199 | Qed. 200 | 201 | (** 202 | No xxxP lemmas allowed, but reflectT and reflectF and case analysis allowed , 203 | *) 204 | Lemma myidP: forall (b : bool), reflect b b. 205 | Proof. 206 | (*D*) move=> b; case: b. 207 | (*D*) exact: ReflectT. 208 | (*D*) exact: ReflectF. 209 | Qed. 210 | 211 | Lemma mynegP: forall (b : bool), reflect (~ b) (~~ b). 212 | Proof. 213 | (*D*) move=> b. case: b. 214 | (*D*) exact: ReflectF. 215 | (*D*) exact: ReflectT. 216 | Qed. 217 | 218 | Lemma myandP: forall (b1 b2 : bool), reflect (b1 /\ b2) (b1 && b2). 219 | Proof. 220 | (*D*) move=> b1 b2. case: b1. 221 | (*D*) case b2. 222 | (*D*) exact: ReflectT. 223 | (*D*) apply: ReflectF. move=> h. by case: h. 224 | (*D*) case b2. 225 | (*D*) apply: ReflectF. move=> h. by case: h. 226 | (*D*) apply: ReflectF. move=> h. by case: h. 227 | Qed. 228 | 229 | Lemma myiffP: 230 | forall (P Q : Prop) (b : bool), 231 | reflect P b -> (P -> Q) -> (Q -> P) -> reflect Q b. 232 | Proof. 233 | (*D*) move=> P Q b Pb PQ QP. move: Pb. case: b. 234 | (*D*) move=> h. case: h. 235 | (*D*) move/PQ=> hp. by apply: ReflectT. 236 | (*D*) move=> hNp. apply: ReflectF. move/QP=> hp. 237 | (*D*) by move/hNp: hp. 238 | (*D*) move=> h. case: h. 239 | (*D*) move/PQ=> hq. by apply: ReflectT. 240 | (*D*) move=> hNp. apply: ReflectF. move/QP=> hp. 241 | (*D*) by move/hNp: hp. 242 | Qed. 243 | 244 | 245 | (** 246 | Some arithmetics 247 | *) 248 | 249 | Fixpoint len (n m : nat) : bool := 250 | match n, m with 251 | | 0 , _ => true 252 | | n'.+1, 0 => false 253 | | n'.+1, m'.+1 => len n' m' 254 | end. 255 | 256 | Lemma lenP: forall n m, reflect (exists k, k + n = m) (len n m). 257 | Proof. 258 | (*D*) move=> n; elim: n. 259 | (*D*) move=> m. apply: (iffP idP). 260 | (*D*) move=> _. by exists m; apply: addn0. 261 | (*D*) by []. 262 | (*D*) move=> n IH m. apply: (iffP idP). 263 | (*D*) case: m. by []. 264 | (*D*) move=> m /= le_nm. move/IH: le_nm=> le_nm. 265 | (*D*) case: le_nm. move=> k eq_xk_m. exists k. 266 | (*D*) by rewrite -eq_xk_m addnS. 267 | (*D*) case: m. 268 | (*D*) case. move=>k. by rewrite addnS. 269 | (*D*) move=> m h. case: h => k. 270 | (*D*) rewrite addnS. move=> eq_kSn_k. case: eq_kSn_k. 271 | (*D*) move=> eq_kn_m. apply/IH. by exists k. 272 | Qed. 273 | 274 | 275 | (* --------------------------------------------------------------------*) 276 | -------------------------------------------------------------------------------- /SummerSchoolSophia/exercise7_solution.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | 3 | 4 | (** *** Exercise 1 : 5 | - Prove this statement by induction using big_nat_recr and big_geq 6 | - alternatively by using big_morph 7 | *) 8 | 9 | Lemma sum_mull f (k n : nat) : 10 | k * (\sum_(0 <= i < n) f i) = \sum_(0 <= i < n) (k * f i). 11 | Proof. 12 | (*D*) (* elim: n => [|n IH]; first by rewrite !big_geq ?muln0. *) 13 | (*D*) (* by rewrite !big_nat_recr //= mulnDr IH. *) 14 | (*D*)by apply: (big_morph (fun n => k * n)) => [x y|]; 15 | [rewrite mulnDr | rewrite muln0]. 16 | Qed. 17 | 18 | (** *** Exercise 2 : 19 | - Prove this statement by using big_morph 20 | *) 21 | 22 | Lemma sum_mulr f (k n : nat) : 23 | (\sum_(0 <= i < n) f i) * k = \sum_(0 <= i < n) (f i * k). 24 | Proof. 25 | (*D*)(* elim: n => [|n IH]; first by rewrite !big_geq. *) 26 | (*D*)(* by rewrite !big_nat_recr //= mulnDl IH. *) 27 | (*D*)by apply: (big_morph (fun n => n * k)) => // x y; rewrite mulnDl. 28 | Qed. 29 | 30 | 31 | (** *** Exercise 3 : 32 | - Prove this statement by induction. 33 | - Relevant lemmas are 34 | - doubleS odd_double addn0 addnn mulnn 35 | - big_mkcond big_nat_recr big_geq 36 | *) 37 | 38 | Lemma sum_odd n : \sum_(0 <= i < n.*2 | odd i) i = n ^ 2. 39 | Proof. 40 | (*D*)elim: n => [|n IHn]; first by rewrite big_geq. 41 | (*D*)rewrite doubleS big_mkcond 2?big_nat_recr // -big_mkcond /=. 42 | (*D*)rewrite {}IHn odd_double /= addn0 -addnn -!mulnn. 43 | (*D*)by rewrite mulSn mulnC mulSn addnA addSn addnC. 44 | Qed. 45 | 46 | 47 | (** *** Exercise 4 : 48 | - Prove this statement by induction. 49 | - Relevant lemmas are 50 | - doubleD muln2 addn2 big_nat_recr big_geq 51 | *) 52 | 53 | Lemma sum_gauss n : (\sum_(0 <= i < n) i).*2 = n * n.-1. 54 | Proof. 55 | (*D*)elim: n => [|n IH]; first by rewrite big_geq. 56 | (*D*)rewrite big_nat_recr //= doubleD {}IH. 57 | (*D*)case: n => [|n /=]; first by rewrite muln0. 58 | (*D*)by rewrite -muln2 -mulnDr addn2 mulnC. 59 | Qed. 60 | 61 | 62 | (** 63 | In what follows we are going to mimic the proof of Gauss : 64 | 65 | << 66 | 1 + 2 + ..... + n.-2 + n.-1 67 | + n.-1 + n.-2 + + 2 + 1 68 | 69 | = n.-1 * n 70 | >> 71 | 72 | **) 73 | 74 | 75 | (** *** Exercise 5 : 76 | - Prove this statement without induction. 77 | - Relevant lemma is big_nat_rev 78 | *) 79 | 80 | Lemma sum_gauss_rev n : \sum_(0 <= i < n) (n.-1 - i) = \sum_(0 <= i < n) i. 81 | Proof. 82 | (*D*)rewrite [RHS]big_nat_rev /=. 83 | (*D*)by case: n => //. 84 | Qed. 85 | 86 | (** *** Exercise 6 : 87 | - Prove this statement without induction. 88 | - Relevant lemma is addnn 89 | *) 90 | Lemma sum_gauss_double n : (\sum_(0 <= i < n) i).*2 = 91 | \sum_(0 <= i < n) i + \sum_(0 <= i < n) (n.-1 - i). 92 | Proof. 93 | (*D*)by rewrite sum_gauss_rev addnn. 94 | Qed. 95 | 96 | 97 | (** *** Exercise 7 : 98 | - Prove this statement without induction. 99 | - Relevant lemma are big_split and eq_big_nat 100 | *) 101 | 102 | Lemma sum_gaussD n : 103 | \sum_(0 <= i < n) i + \sum_(0 <= i < n) (n.-1 - i) = 104 | \sum_(0 <= i < n) n.-1. 105 | Proof. 106 | (*D*)rewrite -big_split /=. 107 | (*D*)apply: eq_big_nat => i Hi. 108 | (*D*)rewrite addnC subnK //. 109 | (*D*)by case: n Hi. 110 | Qed. 111 | 112 | 113 | (** *** Exercise 8 : 114 | - Prove this statement without induction. 115 | - Relevant lemma are sum_nat_const_nat 116 | *) 117 | 118 | Lemma sum_gauss_const n k : \sum_(0 <= i < n) k = n * k. 119 | Proof. 120 | (*D*)by rewrite sum_nat_const_nat subn0. 121 | Qed. 122 | 123 | 124 | (** *** Exercise 9 : 125 | - Prove this statement using exercises 5-8 126 | *) 127 | Lemma sum_gauss_alt1 n : (\sum_(0 <= i < n) i).*2 = n * n.-1. 128 | Proof. 129 | (*D*)by rewrite sum_gauss_double sum_gaussD sum_gauss_const. 130 | Qed. 131 | 132 | 133 | (** *** Exercise 10 : 134 | - Prove this statement directly without using the lemmas 135 | - defined in exercises 5-9 136 | *) 137 | 138 | Lemma sum_gauss_alt2 n : (\sum_(0 <= i < n) i).*2 = n * n.-1. 139 | Proof. 140 | (*D*)rewrite -addnn [X in X + _ = _]big_nat_rev -big_split /=. 141 | (*D*)rewrite -[X in _ = X * _]subn0 -sum_nat_const_nat. 142 | (*D*)apply: eq_big_nat => i. 143 | (*D*)by case: n => // n /andP[iP iLn]; rewrite [_ + _]subnK. 144 | Qed. 145 | 146 | 147 | (** *** Now we try to prove the sum of squares. 148 | 149 | **) 150 | 151 | (** *** We first define the property for a function to be increasing 152 | **) 153 | 154 | 155 | Definition fincr f := forall n, f n <= f n.+1. 156 | 157 | (** *** Exercise 11 : 158 | - Prove this statement by induction 159 | *) 160 | 161 | Lemma fincrD f m n : fincr f -> f m <= f (n + m). 162 | Proof. 163 | (*D*)move=> Hf; elim: n => // n H; exact: leq_trans H (Hf _). 164 | Qed. 165 | 166 | 167 | (** *** Exercise 12 : 168 | - Prove this statement using exercise 11 169 | - Hint : subnK 170 | *) 171 | 172 | Lemma fincr_leq f m n : fincr f -> m <= n -> f m <= f n. 173 | Proof. 174 | (*D*)by move=> Hf Hn; rewrite -(subnK Hn) fincrD. 175 | Qed. 176 | 177 | 178 | (** *** Exercise 13 : 179 | - Proof by induction 180 | - Hints : addnCA subnK fincr_leq big_geq 181 | *) 182 | 183 | Lemma sum_consecutive n f : 184 | fincr f -> f n = \sum_(0 <= i < n) (f i.+1 - f i) + f 0. 185 | Proof. 186 | (*D*)move=> Hf. 187 | (*D*)elim: n => [|n IH]; first by rewrite big_geq. 188 | (*D*)by rewrite big_nat_recr //= addnAC -IH addnC subnK. 189 | Qed. 190 | 191 | 192 | (** *** Exercise 14 : 193 | - Proof using the previous lemma 194 | - Hints : leq_exp2r 195 | *) 196 | Lemma sum_consecutive_cube n : 197 | n^3 = \sum_(0 <= i < n) (i.+1 ^ 3 - i ^ 3). 198 | Proof. 199 | (*D*)rewrite (sum_consecutive _ (fun i => i ^ 3)) ?addn0 //. 200 | (*D*)by move=> m; rewrite leq_exp2r. 201 | Qed. 202 | 203 | 204 | (** *** We give the proof of a technical result 205 | *) 206 | 207 | Ltac sring := 208 | rewrite !(expn1, expnS, =^~mul2n, mulSn, mulnS, addnS, addSn, 209 | mulnDr, mulnDl, add0n, addn0, muln0, addnA, mulnA); 210 | do ! congr (S _); 211 | do ! ((congr (_ + _); [idtac]) || (rewrite [in LHS]addnC ?[in LHS]addnA //)). 212 | 213 | Lemma succ_cube n : n.+1 ^ 3 = n ^ 3 + (3 * n ^ 2 + 3 * n + 1). 214 | Proof. sring. Qed. 215 | 216 | (** *** Exercise 15 : 217 | - Hints : big_split sum_mull sum_gauss sum_gauss_const 218 | *) 219 | Lemma sum_sum3 n : 220 | \sum_(0 <= i < n) (6 * i ^ 2 + 6 * i + 2) = 221 | 6 * (\sum_(0 <= i < n) i ^ 2) + 3 * n * n.-1 + n.*2. 222 | Proof. 223 | (*D*)rewrite big_split big_split /=. 224 | (*D*)rewrite -!sum_mull sum_gauss_const. 225 | (*D*)by rewrite -mulnA -sum_gauss // -mul2n mulnA muln2. 226 | Qed. 227 | 228 | 229 | (** *** Exercise 16 : 230 | - Hints : big_split sum_mull sum_gauss sum_gauss_const 231 | *) 232 | Lemma sum_sum4 n : 233 | (n ^ 3).*2 = 6 * (\sum_(0 <= i < n) i ^ 2) + 3 * n * n.-1 + n.*2. 234 | Proof. 235 | (*D*)rewrite sum_consecutive_cube -sum_sum3 -mul2n sum_mull. 236 | (*D*)apply: eq_big_nat => i Hi. 237 | (*D*)by rewrite succ_cube addKn 2!mulnDr !mulnA. 238 | Qed. 239 | 240 | (** *** Another technical lemma 241 | *) 242 | 243 | Lemma sum_tech n : (n ^ 3).*2 = n * n.-1 * n.-1.*2.+1 + 3 * n * n.-1 + n.*2. 244 | Proof. by case: n => //= n1; sring. Qed. 245 | 246 | 247 | (** *** Exercise 17 : 248 | - Hint : addIn sum_sum4 sum_tech. 249 | *) 250 | Lemma sum_square n : 6 * (\sum_(0 <= i < n) i ^ 2) = n * n.-1 * n.-1.*2.+1. 251 | Proof. 252 | (*D*)apply: (@addIn (3 * n * n.-1 + n.*2)). 253 | (*D*)by rewrite addnA -sum_sum4 sum_tech !addnA. 254 | Qed. 255 | 256 | (** *** Exercise 18 : 257 | - Prove the theorem directly using only sum_gauss and the tactic sring. 258 | *) 259 | Lemma sum_square_alt n : 6 * (\sum_(0 <= i < n) i ^ 2) = n * n.-1 * n.-1.*2.+1. 260 | Proof. 261 | (*D*)apply: (@addIn (3 * n * n.-1 + n.*2)). 262 | (*D*)rewrite addnA -[6]/(2 * 3) -!mulnA mul2n (big_morph _ (mulnDr _) (muln0 _)). 263 | (*D*)rewrite -{1}sum_gauss -doubleMr (big_morph _ (mulnDr _) (muln0 _)). 264 | (*D*)rewrite -{3}[n]muln1 -{3}[n]subn0 -sum_nat_const_nat. 265 | (*D*)rewrite -!doubleD -!big_split /=. 266 | (*D*)rewrite (eq_big_nat _ _ (_ : forall i, _ -> _ = i.+1 ^ 3 - i ^3)) => [|i H]; last first. 267 | (*D*) rewrite -[LHS](addKn (i ^ 3)); congr (_ - _). 268 | (*D*) by sring. 269 | (*D*)apply: etrans (_ : (n^3).*2 = _). 270 | (*D*) congr (_.*2). 271 | (*D*) elim: n => [|n IH]; first by rewrite big_geq. 272 | (*D*) by rewrite big_nat_recr //= IH addnC subnK // leq_exp2r. 273 | (*D*)by case: n => //= n1; sring. 274 | Qed. 275 | -------------------------------------------------------------------------------- /SummerSchoolSophia/lesson3.v: -------------------------------------------------------------------------------- 1 | 2 | From mathcomp Require Import all_ssreflect. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | (** 9 | 10 | ---------------------------------------------------------- 11 | #
# 12 | ** Lesson 3: summary 13 | 14 | - proofs by backchaining/backward reasoning 15 | - proofs by induction 16 | - stack model, tacticials [=>] and [:] 17 | - operations on the stack 18 | 19 | 20 | #


# 21 | 22 | #

# 23 | 24 | ---------------------------------------------------------- 25 | #
# 26 | ** Proofs by backward chaining 27 | 28 | - [move=> names] introduces hypotheses in the context. 29 | - [apply: term] does backward reasoning. 30 | 31 | #
# 32 | *) 33 | Lemma example m p : prime p -> p %| m `! + 1 -> m < p. 34 | Proof. 35 | move=> prime_p. 36 | (* Search "contra". *) 37 | apply: contraLR. 38 | rewrite -leqNgt. 39 | move=> leq_p_m. 40 | rewrite dvdn_addr. 41 | by rewrite gtnNdvd // prime_gt1. 42 | (* Search _ dvdn factorial.*) 43 | apply: dvdn_fact. 44 | by rewrite leq_p_m prime_gt0. 45 | Qed. 46 | (** 47 | #
# 48 | 49 | Remark [dvdn_addr] has a side condition. 50 | 51 | Remark [rewrite] accepts 52 | - many rewrite rules, 53 | - the minus switch [-] to rewrite right to left, 54 | - [//] to solve simple goals, 55 | - [/=] to call the simplification heuristic, 56 | - [/name] to unfold definition [name], 57 | - [{x}] to clear [x] from the context. 58 | 59 | cf #ssreflect documentation on rewrite# 60 | 61 | 62 | Remark [n <= m <= p] is [n <= m && m <= p]. 63 | 64 | #
(notes)
# 65 | 66 | This slide corresponds to 67 | section 2.3.3 of 68 | #the Mathematical Components book# 69 | #
# 70 | 71 | #


# 72 | #

# 73 | 74 | ---------------------------------------------------------- 75 | #
# 76 | ** Proofs by induction 77 | 78 | The [elim:] tactic is like [case:] but gives you 79 | an induction hypothesis. 80 | 81 | #
# 82 | *) 83 | Lemma addn0 m : m + 0 = m. 84 | Proof. 85 | elim: m => [|m IHm]. 86 | by []. 87 | by rewrite addSn IHm. 88 | Qed. 89 | 90 | Lemma muln0 m : m * 0 = 0. 91 | Proof. 92 | elim: m => [|m IHm]. 93 | by []. 94 | by rewrite mulSn IHm. 95 | Qed. 96 | 97 | (** 98 | #
# 99 | 100 | The first pitfall when proving programs by 101 | induction is not generalizing enough the property 102 | being proved before starting the induction. 103 | 104 | #
# 105 | *) 106 | Lemma foldl_cat T R f (z0 : R) (s1 s2 : seq T) : 107 | foldl f z0 (s1 ++ s2) = foldl f (foldl f z0 s1) s2. 108 | Proof. 109 | move: z0. 110 | elim: s1. 111 | by []. 112 | move=> x xs IH. 113 | move=> acc. 114 | rewrite /=. 115 | by rewrite IH. 116 | Qed. 117 | (** 118 | #
# 119 | 120 | #
(notes)
# 121 | 122 | This slide corresponds to 123 | section 2.3.4 of 124 | #the Mathematical Components book# 125 | #
# 126 | #


# 127 | 128 | #

# 129 | 130 | ---------------------------------------------------------- 131 | #
# 132 | ** Goal management 133 | 134 | - naming everything can become bothersome 135 | - but, we should not let the system give random names 136 | - we adopt some sort of "stack & heap" model 137 | 138 | *** The stack model of a goal 139 | #
# 140 | #
#
141 | ci : Ti
142 | …
143 | dj := ej : Tj
144 | …
145 | Fk : Pk ci
146 | …
147 | =================
148 | forall (xl : Tl) …,
149 | let ym := bm in … in
150 | Pn xl -> … ->
151 | Conclusion
152 | #
# 153 | #
# 154 | 155 | #
# 156 | *) 157 | Axiom (Ti Tj Tl : Type) (ej bm : Tj). 158 | Axiom (Pk : Ti -> Type) (Pn : Tl -> Type). 159 | Axiom (Conclusion : Ti -> Tj -> Tj -> Tl -> Type). 160 | 161 | Lemma goal_model_example (ci : Ti) (dj : Tj := ej) (Fk : Pk ci) : 162 | forall (xl : Tl), let ym := bm in Pn xl -> Conclusion ci dj ym xl. 163 | Abort. 164 | (** 165 | #
# 166 | 167 | #
(notes)
# 168 | This slide corresponds to section 169 | #Bookkeeping# of the online docmentation of the ssreflect proof language. 170 | #
# 171 | 172 | #


# 173 | #

# 174 | 175 | ---------------------------------------------------------- 176 | #
# 177 | *** Managing the stack 178 | 179 | - [tactic=> names] executes tactics, pops and names 180 | - [tactic: names] pushes the named objects, then execute tactic 181 | - [move] is the tactic that does nothing (no-op, [idtac]) 182 | 183 | #
# 184 | *) 185 | Lemma goal_model_example (ci : Ti) (dj : Tj := ej) (Fk : Pk ci) : 186 | forall (xl : Tl), let ym := bm in Pn xl -> Conclusion ci dj ym xl. 187 | Proof. 188 | move=> xl ym pnxl. 189 | move: ci Fk. 190 | Abort. 191 | (** 192 | #
# 193 | 194 | #
(notes)
# 195 | #
# 196 | 197 | #


# 198 | #

# 199 | 200 | ---------------------------------------------------------- 201 | #
# 202 | ** intro-pattern and discharge partterns 203 | 204 | You can write 205 | - [tactic=> i_items] where i_item could be a name, or 206 | - [?] name chosen by the system, no user access, 207 | - [_] remove the top of the stack (if possible), 208 | - [//] close trivial subgoals, 209 | - [/=] perform simplifications, 210 | - [//=] do both the previous, 211 | - [->] rewrite using the top of the stack, left to right, 212 | - [<-] same but right to left, 213 | - [{x}] clear name [x] from the context. 214 | - [ [i_items|..|i_items] ] introductions on various sub-goals 215 | when immediately [tactic] is [case] or [elim] 216 | 217 | cf #ssreflect documentation on introduction to the context# 218 | 219 | - [tactic: d_items] where d_item could be a name or a pattern, and 220 | - [tactic] must be [move], [case], [elim], [apply], [exact] or [congr], 221 | - [move: name] clears the name from the context, 222 | - [move: pattern] generalize a subterm of the goal that match the pattern, 223 | - [move: (name)] forces [name] to be a pattern, hence not clearing it. 224 | 225 | cf #ssreflect documentation on discharge# 226 | 227 | #


# 228 | #

# 229 | 230 | ---------------------------------------------------------- 231 | #
# 232 | ** Example 233 | #
# 234 | *) 235 | Lemma goal_model_example (ci : Ti) (dj : Tj := ej) (Fk : Pk ci) : 236 | forall (xl : Tl), let ym := bm in Pn xl -> Conclusion ci dj ym xl. 237 | Proof. 238 | move=> xl /= pnxl. 239 | Fail move=> {xl}. 240 | Fail move=> {dj}. 241 | rewrite /dj {dj}. 242 | move: ci Fk. 243 | move=> {pnxl}. 244 | move=> ci _. 245 | Abort. 246 | (** 247 | #
# 248 | #


# 249 | #

# 250 | ---------------------------------------------------------- 251 | #
# 252 | ** elim and case work on the top of the stack 253 | 254 | [elim: x y z => [t u v | w] ] is the same as 255 | - [move: x y z.] 256 | - [elim.] 257 | - [move=> t u v.] in one sub-goal, [move=> w.] in the other. 258 | 259 | #
# 260 | *) 261 | Lemma addnA m n p : m + (n + p) = (m + n) + p. 262 | Proof. by elim: m => // m IHm; rewrite !addSn IHm. Restart. 263 | Proof. by elim: m => // m; rewrite !addSn => ->. Qed. 264 | 265 | Lemma subnDA m n p : n - (m + p) = (n - m) - p. 266 | Proof. by move: m n; elim=> [//|m IHm]; case. Restart. 267 | Proof. by elim: m n => [|m IHm] []. Qed. 268 | 269 | Lemma andbC : commutative andb. 270 | Proof. move=> b1 b2; case: b1; case: b2. Restart. 271 | Proof. by case; case. Restart. 272 | Proof. by move=> [] []. Qed. 273 | (** 274 | #
# 275 | #


# 276 | #

# 277 | ---------------------------------------------------------- 278 | #
# 279 | ** Example of [foldl_cat] 280 | 281 | #
# 282 | *) 283 | Lemma foldl_cat' T R f (z0 : R) (s1 s2 : seq T) : 284 | foldl f z0 (s1 ++ s2) = foldl f (foldl f z0 s1) s2. 285 | Proof. 286 | move: s1 z0. 287 | elim. 288 | done. 289 | move=> x xs IH. 290 | move=> acc. 291 | rewrite /=. 292 | by rewrite IH. 293 | 294 | Restart. 295 | 296 | Proof. by elim: s1 z0 => [//|x xs IH] acc /=; rewrite IH. Qed. 297 | (** 298 | #
# 299 | 300 | #


# 301 | #

# 302 | 303 | ---------------------------------------------------------- 304 | #
# 305 | ** Lesson 3: sum up 306 | 307 | - [rewrite] takes many arguments 308 | - [apply: t] backward reasonning on the whole goal 309 | - [=>] is pop to context and [:] is push 310 | - [case] case analysis on the top of the stack 311 | - [elim] induction on the top of the stack 312 | 313 | #
# 314 | *) 315 | -------------------------------------------------------------------------------- /Ssreflect15min/quesako.v: -------------------------------------------------------------------------------- 1 | (*** SSReflect, an extension of Coq's tactic language ***) 2 | 3 | (** History **) 4 | 5 | 6 | (* 7 | 8 | - 2006 : First version, part of the Coq proof of the Four Colour Theorem 9 | - 2008 : First standalone release of Ssreflect (v1.1), for Coq v8.1 10 | - 2009 : First version of the documentation manual 11 | - 2012 : Completion of the proof of the Odd Order Theorem 12 | - 2017 : Integration in the standard distribution of Coq 13 | - today : used in various coq developments, of various nature 14 | (analysis, quickchick, elliptic curves, coqcombi, etc.) 15 | *) 16 | 17 | 18 | (** Name **) 19 | 20 | (* SSReflect is coined after "Small Scale Reflection", a formalization 21 | methodology making the most of the status of computation in TT. 22 | 23 | 24 | Ssreflect may thus refer to three different things: a formalization methodology, 25 | a tactic language, and a corpus of formalized libraries, based on the former 26 | methodology and developped using the latter language. A more appropriate name 27 | for the libraries is "Mathematical Components". 28 | 29 | In this file, we give a describe the small scale reflection method in a nutshell, 30 | so as to illustrate the motivations for some features of the tactic language. 31 | 32 | We start by writing proof scripts that do not depend on the ssreflect extension 33 | of the tactic language, because the methodology is independent. 34 | *) 35 | 36 | 37 | (* == Reminder: large scale reflection *) 38 | 39 | From Coq Require Import ZArith. 40 | 41 | Open Scope Z_scope. 42 | 43 | Definition c : Z := 44 | 90377629292003121684002147101760858109247336549001090677693. 45 | 46 | 47 | (* An oracle produces these numbers... *) 48 | Definition f1 := 588120598053661. 49 | Definition f2 := 260938498861057. 50 | Definition f3 := 760926063870977. 51 | Definition f4 := 773951836515617. 52 | 53 | (* We can check this identity using Coq: *) 54 | Lemma factors_c : c = f1 * f2 * f3 * f4. 55 | (* the proof goes "by computation", because '_ * _' is a program *) 56 | Proof. Time reflexivity. Qed. 57 | 58 | (* Computational proofs on moderate size objects, like in the above example 59 | already require thinking a bit about data-structures and complexity. *) 60 | Close Scope Z_scope. 61 | 62 | (* E.g. unary numbes are not an option. *) 63 | (* Time Eval compute in (1000 * 1000 : nat). *) 64 | 65 | (* A reflection-based tactic replace explicit deduction steps, 66 | logged in proof terms, with computational steps, checked by the kernel's 67 | reduction machine. Reflection uses an inductive type (the "reflected" one) 68 | as an abstract, symbolic, data strutcure which can be used to write programs. 69 | The tactic checks that these computations can be interpreted as a valid 70 | propositional equality on a source type. For instance, the ring tactic reflects 71 | a type equipped with a, possibly axiomatic and non-computational, 72 | ring structure into a type for polynomial expressions. *) 73 | 74 | 75 | Require Import Arith Ring. 76 | 77 | Goal forall (n : nat) (b : bool) (f : bool -> nat), 78 | n + (f b) * n + n * (n + 1) = n * (n + (f b)) + 2 * n. 79 | Proof. 80 | intros. ring. 81 | Qed. 82 | 83 | 84 | (* == Small scale reflection *) 85 | 86 | (* In this case as well, the central idea is to replace deductive steps 87 | by computation. But at a smaller and pervasive scale. *) 88 | 89 | 90 | (* Standard library's definition of <= on natural numbers: an 91 | inductive representation of proofs. *) 92 | 93 | Print le. 94 | 95 | 96 | (* A small scale reflection version: a boolean program. For the sake of 97 | modularity, we expess it using the pseudo-substraction on natural 98 | numbers, combined with a boolean equality test. *) 99 | 100 | Fixpoint subn (n m : nat) : nat := 101 | match n, m with 102 | |_, 0 => n 103 | |0, _ => n 104 | |S n, S m => subn n m 105 | end. 106 | 107 | Fixpoint eqn (n m : nat) : bool := 108 | match n, m with 109 | |0, 0 => true 110 | |S n, S m => eqn n m 111 | |_ , _ => false 112 | end. 113 | 114 | Definition leq (n m : nat) : bool := eqn (subn n m) 0. 115 | 116 | (* Let us compare a proof that 16 <= 64, in both approaches *) 117 | Goal le 16 64. 118 | Proof. 119 | auto. 120 | auto 50. 121 | Show Proof. 122 | Qed. 123 | 124 | Goal leq 16 64 = true. 125 | auto. 126 | Show Proof. 127 | Qed. 128 | 129 | (* Small scale reflection at work: *) 130 | Goal forall n m, leq n m = true -> leq (S n) (S m) = true. 131 | Proof. 132 | intros n m h. 133 | (* Just like the proof of factor_c was by "reflexivity", the 134 | proof of the current goal is "exactly" h, because the left-hand 135 | sides of the equalities are convertible. *) 136 | apply h. 137 | Qed. 138 | 139 | (* The same proof with the le definition require an explicit 140 | induction step *) 141 | Goal forall n m, le n m -> le (S n) (S m). 142 | Proof. 143 | intros n m h. 144 | induction h as [| m ihnm ihSS]. 145 | easy. 146 | constructor. 147 | easy. 148 | Qed. 149 | 150 | (** Selected basic features of the ssreflect tactic language **) 151 | 152 | (* From now on, we will use the tactic extension. It is distributed 153 | as part of the Coq proof assistant, and loaded by the following Import *) 154 | From Coq Require Import ssreflect. 155 | 156 | (* The extension implements an improved rewrite tactic language. We 157 | mention only two important features: 158 | == a diferent matching algorithm, rigid on head symbols and using conversion 159 | for the internal subterms. 160 | == a pattern selection mechanism based on this matching. *) 161 | 162 | Goal forall n m c : nat, (forall a b, a + b = c) -> 163 | (n + m) + m + (n + m) = (n + m) + n + (n + m). 164 | Proof. 165 | intros n m c h. 166 | rewrite [n + _]h. 167 | Admitted. 168 | 169 | 170 | (* The extension implements a view mechanism, to combine deductive 171 | steps (e.g. using an implication), with bureaucracy ones (e.g. 172 | introduction/generalization, case analysis). This is quite useful 173 | to manage going back and forth between bool and Prop. Note: bool can be seen 174 | as a type reflecting (certain) inhabitants of Prop. 175 | *) 176 | 177 | From mathcomp Require Import ssrfun ssrbool. 178 | 179 | Goal forall b1 b2 b3 : bool, b1 /\ b2 -> b1 && b2 || b3. 180 | intros b1 b2 b3. 181 | move/andP. 182 | case/andP. 183 | Admitted. 184 | 185 | 186 | (* Another strong motivation behind the design choices applied in the 187 | ssreflect tactic langage is to ease the maintenance/upkeep of formal libraries. 188 | Features associated with this objective are: 189 | == fewer tactics: case, apply, elim, rewrite (move) 190 | == two main tacticals and many switches 191 | == clearing policy for an easier book-keeping 192 | == enhanced support for forward chaining: have, suff, wlog 193 | *) 194 | 195 | 196 | (** SSReflect is only a stepping stone to build libraries. 197 | Demo: Prove that the concatenation (s1 ++ s2) of two lists is 198 | duplicate-free if and only if (s2 ++ s2) is also duplicate-free. **) 199 | 200 | From mathcomp Require Import all_ssreflect. 201 | 202 | Module ExerciseSeq. 203 | 204 | (* Implement a unary predicate on lists characterizing 205 | "duplicate free lists" *) 206 | 207 | (* The standard library answer: 208 | Inductive NoDup (A : Type) : list A -> Prop := 209 | NoDup_nil : NoDup nil 210 | | NoDup_cons : forall (x : A) (l : list A), 211 | ~ In x l -> NoDup l -> NoDup (x :: l) 212 | *) 213 | 214 | 215 | (* Compcert's take: 216 | Inductive list_norepet (A: Type) : list A -> Prop := 217 | | list_norepet_nil: list_norepet nil 218 | | list_norepet_cons: forall hd tl, 219 | ~(In hd tl) -> list_norepet tl -> list_norepet (hd :: tl). 220 | *) 221 | 222 | 223 | (* Mathematical Components's take: *) 224 | 225 | 226 | Section ListNorepet. 227 | 228 | (* eqType packages: 229 | T : Type 230 | _ == _ : T -> T -> bool 231 | eqP : forall x y, x == y <-> x = y 232 | *) 233 | 234 | Variable T : eqType. 235 | Implicit Type s : list T. 236 | 237 | (* Lists on an eqType have a boolean membership unary predicate, 238 | with a generic infix notation *) 239 | 240 | Variables (foolist : list T) (foox : T). 241 | 242 | Check (mem foolist). 243 | Check foox \in foolist. 244 | Check foox \notin foolist. 245 | 246 | Check forall s x, x \in x :: s (* = true *). 247 | 248 | (* Now implement uniq *) 249 | Fixpoint uniq s : bool := 250 | if s is x :: s' then (x \notin s') && uniq s' else true. 251 | 252 | Fact uniq_nil : uniq nil. Proof. done. Qed. 253 | 254 | Fact uniq_cons x s : x \notin s -> uniq s -> uniq (x :: s). 255 | Proof. 256 | by move=> nsx us; rewrite /= nsx. 257 | Qed. 258 | 259 | Print has. 260 | 261 | Lemma cat_uniq s1 s2 : 262 | uniq (s1 ++ s2) = [&& uniq s1, ~~ has (mem s1) s2 & uniq s2]. 263 | Proof. 264 | elim: s1 => [| x s1 ihs1] /=; first by rewrite has_pred0 //. 265 | rewrite mem_cat negb_or. rewrite has_sym. rewrite /= negb_or. 266 | rewrite -!andbA. do !bool_congr. rewrite has_sym. done. 267 | Qed. 268 | 269 | Fact uniq_cat s1 s2 : uniq (s1 ++ s2) = uniq (s2 ++ s1). 270 | Proof. 271 | rewrite !cat_uniq has_sym. do !bool_congr. 272 | Qed. 273 | 274 | End ListNorepet. 275 | 276 | End ExerciseSeq. 277 | 278 | 279 | -------------------------------------------------------------------------------- /AnIntroductionToSmallScaleReflectionInCoq/section3.v: -------------------------------------------------------------------------------- 1 | (******************************************************************************) 2 | (* Solutions of exercises : A script language for structured proofs *) 3 | (******************************************************************************) 4 | 5 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat div seq. 6 | From mathcomp Require Import path choice fintype tuple finset. 7 | 8 | Set Implicit Arguments. 9 | Unset Strict Implicit. 10 | Import Prenex Implicits. 11 | 12 | 13 | (******************************************************************************) 14 | (* Exercise 3.1.1 *) 15 | (******************************************************************************) 16 | 17 | Section Tauto. 18 | 19 | Variables A B C : Prop. 20 | 21 | (* The exact tactic takes its argument on top of the goal stack *) 22 | Lemma tauto1 : A -> A. 23 | Proof. 24 | exact. 25 | Qed. 26 | 27 | (* exact: hAB behaves like (by apply: hAB) and hence finds the hypothesis A 28 | in the context needed to solve the goal *) 29 | Lemma tauto2 : (A -> B) -> (B -> C) -> A -> C. 30 | Proof. 31 | move=> hAB hBC hA. 32 | apply: hBC. 33 | exact: hAB. 34 | Qed. 35 | 36 | Lemma tauto4 : A /\ B <-> B /\ A. 37 | Proof. by split; case=> h1 h2; split. Qed. 38 | 39 | End Tauto. 40 | 41 | (******************************************************************************) 42 | (* Exercise 3.1.2 *) 43 | (******************************************************************************) 44 | 45 | Section MoreBasics. 46 | 47 | Variables A B C : Prop. 48 | Variable P : nat -> Prop. 49 | 50 | Lemma foo1 : ~(exists x, P x) -> forall x, ~P x. 51 | move=> h x Px. 52 | apply: h. 53 | by exists x. 54 | Qed. 55 | 56 | 57 | Lemma foo2 : (exists x, A -> P x) -> (forall x, ~P x) -> ~A. 58 | Proof. 59 | case=> x hx hP hA. 60 | apply: (hP x). 61 | exact: hx. 62 | Qed. 63 | End MoreBasics. 64 | 65 | (******************************************************************************) 66 | (* Exercise 3.1.3 *) 67 | (******************************************************************************) 68 | 69 | (* Try Search "<=" to see the various notations featuring <= *) 70 | 71 | (* The first line of the returned answer gives the name of 72 | the predicate (leq) *) 73 | Search "_ <= _". 74 | 75 | (* The Print command shows that leq is defined using subtraction *) 76 | Print leq. 77 | 78 | (* This time we see that the first line of the answer gives the *) 79 | (*answer: there is no constant defining "<" but the notation hides a *) 80 | (* defintiion using leq *) 81 | 82 | Search "_ < _". 83 | 84 | (* Both cases gereated by the induction lead to trivial goals *) 85 | Lemma tuto_subnn : forall n : nat, n - n = 0. 86 | Proof. by elim=> [|n ihn]. Qed. 87 | 88 | (* This proof is an induction on m, followed by a case analysis on *) 89 | (* the second n. Base case is trivial (hence discarded by the // *) 90 | (* switch. *) 91 | Lemma tuto_subn_gt0 : forall m n, (0 < n - m) = (m < n). 92 | Proof. elim=> [|m IHm] [|n] //; exact: IHm. Qed. 93 | 94 | 95 | (* This proof starts the same way as the one of subn_gt0. *) 96 | (* Then two rewriting are chained to reach to trivial subgoals *) 97 | Lemma tuto_subnKC : forall m n : nat, m <= n -> m + (n - m) = n. 98 | Proof. 99 | elim=> [|m IHm] [|n] // Hmn. 100 | Search _ (_.+1 + _ = _.+1). 101 | by rewrite addSn IHm. 102 | Qed. 103 | 104 | (* The first Search command suggests we need to use something like *) 105 | (*subn_add2r, hence to transform n into n - p + p *) 106 | (* The second Search command finds the appropriate lemma. We rewrite *) 107 | (*it from right to left, only at the second occurrence of n, the *) 108 | (*condition of the lemma is fullfilled thanks to the le_pn hypothesis, *) 109 | (*hence the generated subgoal is closed by the // switch *) 110 | 111 | Lemma tuto_subn_subA : forall m n p, p <= n -> m - (n - p) = m + p - n. 112 | Proof. 113 | move=> m n p le_pn. 114 | Search _ (_ + _ - _ = _) in ssrnat. 115 | Search _ (_ - _ + _). 116 | rewrite -{2}(subnK le_pn) // subnDr. 117 | done. 118 | Qed. 119 | 120 | (******************************************************************************) 121 | (* Exercise 3.5.1 *) 122 | (******************************************************************************) 123 | 124 | (* The Check instruction only gives the type of a constant, the *) 125 | (*statement of a lemma. The Print command gives the body of the *) 126 | (*definition, and possibly some extra information (scope, implicit *) 127 | (*arguments,...). Print should not be used in general on lemmas *) 128 | (* since the body of a proof is seldom relevant...*) 129 | 130 | Print edivn. 131 | Print edivn_rec. 132 | Print edivn_spec. 133 | (* The edivn_spec is defined as a CoInductive predicate. The intended *) 134 | (* meaning is not to define an coinductive structure, but rather an *) 135 | (* inductive one. CoInductive in this case indeed behaves as Inductive *) 136 | (*but does not generate an induction principle, which would be is *) 137 | (*useless in this case. *) 138 | 139 | 140 | (******************************************************************************) 141 | (* Exercise 3.5.2 *) 142 | (******************************************************************************) 143 | 144 | (* At this point, the top of the goal stack was featuring three *) 145 | (*natural numers: an induction on the first one generated two subgoals. *) 146 | (*In the second subgoal, corresponding to the inductive case, the *) 147 | (*generated natural number and induction hypothesis have been *) 148 | (*introduced by the branching intro pattern [| n IHn], which leaves *) 149 | (*the first subgoal, corresponding to the base case of the induction, *) 150 | (*unchanged. Then [|m] performs in both subgoals a case analysis on the *) 151 | (*second natural number. This case analysis again leads to two new *) 152 | (*subgoal for each initial branch (which makes four subgoals). In the *) 153 | (*second case of the analysis, the new natural number is introduced *) 154 | (*and named m. Then the third natural number is uniformly introdued in *) 155 | (*the four cases under the name q. Finally the //= switch simplifies *) 156 | (*in the four bracnhes and closes the goals that have become trivial *) 157 | (*(i.e. which are solved by done). *) 158 | 159 | (******************************************************************************) 160 | (* Exercise 3.5.3 *) 161 | (******************************************************************************) 162 | 163 | (* The pattern [// | le_dm] closes the first subgoal and introduces an *) 164 | (*hypothesis named le_dm in the second subgoal. This is equivalent to *) 165 | (* // le_dm. *) 166 | 167 | (******************************************************************************) 168 | (* Exercise 3.5.4 *) 169 | (******************************************************************************) 170 | 171 | (* Replacing (ltnP m d) by ltnP does not change the behaviour of the *) 172 | (*script: Coq's unification is powerful enough to guess the arguments *) 173 | (*in this case since there is only one instance of the comparison in *) 174 | (* (_ <_) the goal. Arguments are mandatory only in the case the frist *) 175 | (*occurrence of the comparison is not the one the user whould like to *) 176 | (*pick as support for ase analysis. *) 177 | 178 | Check ltnP. 179 | Print ltn_xor_geq. 180 | 181 | (* For any two natural numbers n and m, (ltn_xor_geq m n) is a binary *) 182 | (*relation on boolean. Its inductive definition has two constructors *) 183 | (*and states that *) 184 | (* - (false, true) is in the relation(ltn_xor_geq m n) as soon as 185 | m < n *) 186 | (* - (true, false) is in the relation(ltn_xor_geq m n) as soon as 187 | n <= m *) 188 | (* The inductive construction implies that these two rules are the *) 189 | (*only ways to populate the relation.*) 190 | (* Now the theorem:*) 191 | (* ltnP : forall m n : nat, ltn_xor_geq m n (n <= m)(m < n)*) 192 | (*proves that for any two natural numbers m and n, there are only two*) 193 | (*possible situations:*) 194 | (* - either m < n (first rule of the relation definition), and in this *) 195 | (*case n <= m = false and m < n = true *) 196 | (* - or n <= m (second rule of the relation definition, and in this *) 197 | (*case n <= m = true and m < n = false *) 198 | (* A case analysis on this result hence generates two subgoals, one *) 199 | (*for each constructor of ltn_xor_gep. In each subgoal the hypothesis *) 200 | (*of the rule (repectively m < n and n <= m) appears on the stack. Also *) 201 | (*every occurrence of (n <= m) and (m < n) is replaced by the value *) 202 | (*imposed by the ltn_xor_gep constructor used in the branch.*) 203 | 204 | CoInductive tuto_compare_nat (m n : nat) : bool -> bool -> bool -> Set := 205 | | TCompareNatLt of m < n : tuto_compare_nat m n true false false 206 | | TCompareNatGt of m > n : tuto_compare_nat m n false true false 207 | | TCompareNatEq of m = n : tuto_compare_nat m n false false true. 208 | 209 | (* Let's check against what is defined in the ssrnat library *) 210 | (* CHANGE : min and max have been added *) 211 | Lemma tuto_ltngtP : forall m n, 212 | compare_nat m n (minn n m) (minn m n) (maxn n m) 213 | (maxn m n) (n == m) (m == n) (n <= m) (m <= n) 214 | (n < m) (m < n). 215 | Proof. 216 | move=> m n. 217 | rewrite !ltn_neqAle [_ == n]eq_sym; have [mn|] := ltnP m n. 218 | by rewrite ltnW // gtn_eqF //; constructor. 219 | rewrite leq_eqVlt; case: ltnP; rewrite ?(orbT, orbF) => //= lt_nm eq_nm. 220 | by rewrite ltn_eqF //; constructor. 221 | by rewrite eq_nm (eqP eq_nm); constructor. 222 | Qed. 223 | -------------------------------------------------------------------------------- /SummerSchoolSophia/lesson7.v: -------------------------------------------------------------------------------- 1 | 2 | From mathcomp Require Import all_ssreflect. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | (** 9 | 10 | ---------------------------------------------------------- 11 | #
# 12 | ** Lesson 7: summary 13 | 14 | - generic notations and theories 15 | - interfaces 16 | - parametrizing theories 17 | - the BigOp library (the theories of fold) 18 | 19 | Let's start with a lie and then make it true: 20 | 21 | #
# 22 | Coq is an object oriented 23 | programming language. 24 | #
# 25 | 26 | 27 | #
# 28 | 29 | 30 | ---------------------------------------------------------- 31 | #
# 32 | ** Generic notations and theories 33 | 34 | Polymorphism != overloading. 35 | 36 | Example: the [==] computable equality 37 | 38 | #
# 39 | *) 40 | 41 | 42 | (** 43 | #
# 44 | 45 | Polymorphism 46 | 47 | #
# 48 | 49 | *) 50 | 51 | Check (_ = _). 52 | 53 | Check true = false. 54 | 55 | Check (eq true false). 56 | 57 | Check @eq. 58 | 59 | Check (@eq _ true false). 60 | 61 | Check (@eq bool true false). 62 | 63 | 64 | (** 65 | #
# 66 | 67 | Overloading : looking inside types 68 | 69 | #
# 70 | 71 | *) 72 | 73 | Check (_ == _). 74 | 75 | Check true == false. 76 | 77 | Check (@eq_op _ true false). 78 | 79 | Check (@eq_op (Equality.clone _ bool) true false). 80 | 81 | Check 3 == 4. 82 | 83 | Check [::] == [:: 2; 3; 4]. 84 | 85 | Section T. 86 | 87 | Variable T : eqType. 88 | Variable x : T. 89 | 90 | Eval lazy in x == x. 91 | 92 | End T. 93 | 94 | (** 95 | #
# 96 | 97 | Object oriented flavor 98 | 99 | #
# 100 | 101 | *) 102 | 103 | Eval lazy in true == false. 104 | 105 | Eval lazy in 3 == 4. 106 | 107 | Eval lazy in [::] == [:: 2; 3; 4]. 108 | 109 | Check (3, true) == (4, false). 110 | 111 | Eval lazy in (3, true) == (4, false). 112 | 113 | (** 114 | #
# 115 | 116 | Overloading may fail, polymorphism never 117 | 118 | #
# 119 | 120 | *) 121 | 122 | Fail Check (fun x => x) == (fun y => y). 123 | 124 | Check (fun x => x) = (fun y => y). 125 | 126 | (** 127 | #
# 128 | 129 | Type inference 130 | 131 | #
# 132 | 133 | *) 134 | 135 | Check [eqType of bool]. 136 | 137 | Fail Check [eqType of bool -> bool]. 138 | 139 | Check [eqType of {ffun bool -> bool}]. 140 | 141 | Check [eqType of nat]. 142 | 143 | Fail Check [eqType of nat -> nat]. 144 | 145 | Check [eqType of {ffun 'I_256 -> nat}]. 146 | 147 | Check [eqType of seq nat]. 148 | 149 | Fail Check [eqType of seq (nat -> nat)]. 150 | 151 | Check [eqType of seq {ffun 'I_256 -> nat}]. 152 | 153 | (** 154 | #
# 155 | 156 | We call [eqType] an interface. With some "approximation" 157 | [eqType] is defined as follows: 158 | 159 | << 160 | 161 | Module Equality. 162 | 163 | Structure type : Type := Pack { 164 | sort : Type; 165 | op : sort -> sort -> bool; 166 | axiom : ∀x y, reflect (x = y) (op x y) 167 | }. 168 | 169 | 170 | End Equality 171 | >> 172 | 173 | #
(notes)
# 174 | This slide corresponds to 175 | section 5.4 of 176 | #the Mathematical Components book# 177 | #
# 178 | #


# 179 | #

# 180 | 181 | 182 | ---------------------------------------------------------- 183 | #
# 184 | ** Interfaces 185 | 186 | Mathematical Components defines a hierarchy 187 | of interfaces. They group notations and 188 | theorems. 189 | 190 | # # 191 | 192 | Let's use the theory of [eqType] 193 | 194 | #
# 195 | *) 196 | 197 | About eqxx. 198 | 199 | About eq_refl. 200 | 201 | Lemma test_eq (*(T : eqType) (x : T)*) : 202 | (3 == 3) && (true == true) (*&& (x == x)*). 203 | Proof. 204 | rewrite eqxx. 205 | rewrite eqxx. 206 | (* rewrite eqxx. *) 207 | by []. 208 | Qed. 209 | 210 | (** 211 | #
# 212 | 213 | Interfaces do apply to registered, concrete examples 214 | such as [bool] or [nat]. They can also apply to variables, 215 | as long as their type is "rich" ([eqType] is richer than [Type]). 216 | 217 | #
(notes)
# 218 | This slide corresponds to 219 | section 5.5 and 7 of 220 | #the Mathematical Components book# 221 | #
# 222 | 223 | #


# 224 | #

# 225 | 226 | 227 | ---------------------------------------------------------- 228 | #
# 229 | ** Theories over an interface 230 | 231 | Interfaces can be used to parametrize an 232 | entire theory 233 | 234 | #
# 235 | *) 236 | Module Seq. Section Theory. 237 | 238 | Variable T : eqType. 239 | 240 | Implicit Type s : seq T. 241 | 242 | Fixpoint mem_seq s x := 243 | if s is y :: s1 244 | then (y == x) || mem_seq s1 x 245 | else false. 246 | 247 | (* the infix \in and \notin are generic, not 248 | just for sequences. *) 249 | 250 | Fixpoint uniq s := 251 | if s is x :: s1 252 | then (x \notin s1) && uniq s1 253 | else true. 254 | 255 | Fixpoint undup s := 256 | if s is x :: s1 then 257 | if x \in s1 then undup s1 else x :: undup s1 258 | else [::]. 259 | 260 | End Theory. End Seq. 261 | 262 | About undup_uniq. 263 | 264 | Eval lazy in (undup [::1;3;1;4]). 265 | 266 | Lemma test : uniq (undup [::1;3;1;4]). 267 | Proof. 268 | by rewrite undup_uniq. 269 | Qed. 270 | 271 | (** 272 | #
# 273 | 274 | Others interfaces 275 | 276 | **) 277 | 278 | Section Interfaces. 279 | 280 | Variable chT : choiceType. 281 | 282 | Check (@sigW chT). 283 | 284 | Check [eqType of chT]. 285 | 286 | Variable coT : countType. 287 | 288 | Check [countType of nat]. 289 | Check [choiceType of coT]. 290 | Check [choiceType of nat * nat]. 291 | Check [choiceType of seq coT]. 292 | 293 | Variable fT : finType. 294 | 295 | Check [finType of bool]. 296 | Check [finType of 'I_10]. 297 | Check [finType of {ffun 'I_10 -> fT}]. 298 | Check [finType of bool * bool]. 299 | 300 | End Interfaces. 301 | 302 | 303 | (** 304 | #
# 305 | 306 | #
(notes)
# 307 | This slide corresponds to 308 | section 5.6 of 309 | #the Mathematical Components book# 310 | #
# 311 | 312 | #


# 313 | #

# 314 | 315 | ---------------------------------------------------------- 316 | #
# 317 | ** Generic theories: the BigOp library 318 | 319 | The BigOp library is the canonical example 320 | of a generic theory. It it about the 321 | [fold] iterator we studied in lesson 1, 322 | and the many uses it can have. 323 | 324 | #
# 325 | *) 326 | 327 | Lemma sum_odd_3 : 328 | \sum_(0 <= i < 6 | odd i) i = 3^2. 329 | Proof. 330 | rewrite unlock /=. 331 | by []. 332 | Qed. 333 | 334 | About big_mkcond. 335 | About big_nat_recr. 336 | 337 | Lemma sum_odd_3_bis : 338 | \sum_(0 <= i < 6 | odd i) i = 3^2. 339 | Proof. 340 | rewrite big_mkcond big_nat_recr //= -big_mkcond /=. 341 | Abort. 342 | 343 | Lemma prod_odd_3_bis : (* try [maxn/0] and also [maxn/1] *) 344 | \big[muln/1]_(0 <= i < 6 | odd i) i = 3^2. 345 | Proof. 346 | rewrite big_mkcond big_nat_recr //= -big_mkcond /=. 347 | Abort. 348 | 349 | (** 350 | #
# 351 | 352 | Most of the lemmas require the operation to be a monoid, 353 | some others to be a commutative monoid. 354 | 355 | #
# 356 | *) 357 | 358 | About bigD1. 359 | 360 | (** 361 | #
# 362 | 363 | Searching for bigop 364 | 365 | #
# 366 | *) 367 | 368 | Lemma sum_odd_even_all n : 369 | \sum_(0 <= i < n) i = 370 | \sum_(0 <= i < n | odd i) i + \sum_(0 <= i < n | ~~ odd i) i. 371 | Proof. 372 | Search _ (~~ _) inside bigop. 373 | by rewrite (bigID odd). 374 | Qed. 375 | 376 | (** 377 | #
# 378 | 379 | Primer for bigop 380 | 381 | #
# 382 | *) 383 | 384 | Section Primer. 385 | 386 | Variable n: nat. 387 | Variable f : 'I_n -> nat. 388 | Variable g : nat -> nat. 389 | 390 | Check \big[addn/0]_(i <- [::1; 4; 5] | odd i) g i. 391 | 392 | Check \big[addn/0]_(i : 'I_n | odd i) f i. 393 | 394 | Definition oddIn := [pred i | odd (i : 'I_n)]. 395 | 396 | Check \big[addn/0]_(i in oddIn) f i. 397 | 398 | Goal \sum_(i in oddIn) i = \sum_(i < n | odd i) i. 399 | Proof. 400 | by []. 401 | Qed. 402 | 403 | Check \big[addn/0]_(0 <= i < n | odd i) g i. 404 | 405 | Goal \sum_(0 <= i < n | odd i) g i = \sum_(i < n | odd i) g i. 406 | Proof. 407 | Check big_mkord. 408 | rewrite big_mkord. 409 | by []. 410 | Qed. 411 | 412 | Goal \sum_(0 <= i < n |odd i) i.*2 = \sum_(0 <= i < n|odd i) (i + i). 413 | Proof. 414 | Fail rewrite addnn. 415 | About eq_bigr. 416 | apply: eq_bigr. 417 | move=> i Hi. 418 | by rewrite addnn. 419 | Qed. 420 | 421 | Goal \sum_(0 <= i < n |odd i) i.*2 = \sum_(0 <= i < n | odd i) (i + i). 422 | Proof. 423 | About eq_bigr. 424 | rewrite (eq_bigr (fun i => i + i)) //. 425 | move=> i Hi. 426 | by rewrite addnn. 427 | Qed. 428 | 429 | Goal (\sum_(i < n|odd i) i).*2 = \sum_(i < n |odd i) i.*2. 430 | Proof. 431 | About big_morph. 432 | Fail rewrite big_morph. 433 | rewrite (big_morph _ (_ : {morph double : x y / x + y}) (_ : 0.*2 = 0)). 434 | - by []. 435 | - move=> x y; exact: doubleD. 436 | by []. 437 | Qed. 438 | 439 | End Primer. 440 | 441 | 442 | (** 443 | #
# 444 | 445 | #
(notes)
# 446 | This slide corresponds to 447 | section 5.7 of 448 | #the Mathematical Components book# 449 | #
# 450 | 451 | 452 | #


# 453 | #

# 454 | 455 | ---------------------------------------------------------- 456 | #
# 457 | ** Sum up 458 | 459 | - Coq is an object oriented language ;-) 460 | 461 | - in the Mathematical Components library [xxType] is an 462 | interface (eg [eqType] for types with an equality test). 463 | Notations and theorems are linked to interfaces. 464 | Interfaces are organized in hierarchies (we just saw a picture, 465 | how it works can be found in the book). 466 | 467 | #
# 468 | 469 | *) 470 | -------------------------------------------------------------------------------- /SummerSchoolSophia/lesson2.v: -------------------------------------------------------------------------------- 1 | 2 | From mathcomp Require Import all_ssreflect. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | (** 9 | 10 | ---------------------------------------------------------- 11 | #
# 12 | ** Lesson 2: summary 13 | 14 | - statements 15 | - proofs by computation 16 | - proofs by case split 17 | - proofs by rewriting 18 | 19 | #


# 20 | 21 | #

# 22 | 23 | ---------------------------------------------------------- 24 | #
# 25 | ** Formal proofs 26 | 27 | Today we learn how to state and prove theorems. 28 | We don't do that in the void, nor without a methodology. 29 | 30 | We work on top of the Mathematical Components library 31 | and we follow the Small Scale Reflection approach using 32 | the SSReflect proof language. 33 | 34 | The Mathematical Components library can be 35 | #browsed online#. 36 | The modules of interest are 37 | #ssrnat# 38 | and 39 | #ssrbool# 40 | (see the headers for the doc). 41 | 42 | The SSReflect proof language 43 | (#reference manual#) 44 | is covered in full details by the Mathematical Components book. 45 | Here we just cover the basics. 46 | 47 | 48 | #


# 49 | #

# 50 | 51 | ---------------------------------------------------------- 52 | #
# 53 | ** Formal statements 54 | 55 | Most of the statements that we consider in Mathematical 56 | Components are equalities. 57 | 58 | #
# 59 | *) 60 | 61 | Check (_ = _). 62 | Locate "_ = _". 63 | 64 | (** 65 | #
# 66 | 67 | For example, we can equate two expressions representing natural numbers. 68 | 69 | #
# 70 | 71 | *) 72 | 73 | Lemma addnA: forall (m n k : nat), m + (n + k) = (m + n) + k. 74 | Abort. 75 | 76 | (** 77 | #
# 78 | 79 | Addition is defined as left-associative. 80 | 81 | #
# 82 | *) 83 | 84 | Lemma addnA: forall (m n k : nat), m + (n + k) = m + n + k. 85 | Abort. 86 | 87 | (** 88 | #
# 89 | 90 | Quantifications can be set as parameters before the colon. 91 | 92 | #
# 93 | *) 94 | 95 | Lemma addnA (m n k : nat) : m + (n + k) = m + n + k. 96 | Abort. 97 | 98 | (** 99 | #
# 100 | 101 | In lesson 1 we have defined many boolean tests that can 102 | play the role of (decidable) predicates. 103 | 104 | #
# 105 | *) 106 | 107 | Check 0 <= 4. (* not a statement *) 108 | Check (0 <= 4) = true. (* a statement we can prove *) 109 | 110 | (** 111 | #
# 112 | 113 | #
# 114 | Motto: whenever possible predicates are expressed as a programs. 115 | #
# 116 | 117 | This choice has a deep impact on the proofs we make in lesson 2 and 3 and 118 | the way we can form new types in lesson 4. 119 | 120 | Booleans can be coerced into statements. 121 | #
# 122 | *) 123 | 124 | Check is_true (* Definition is_true b := b = true *). 125 | 126 | (** 127 | #
# 128 | 129 | Tests can be turned into statements. 130 | 131 | #
# 132 | *) 133 | 134 | Check (_ <= _). 135 | 136 | Check is_true (_ <= _). 137 | 138 | Lemma leq0n (n : nat) : is_true (0 <= n). 139 | Abort. 140 | 141 | (** 142 | #
# 143 | 144 | the [is_true] 145 | "coercion" is automatically inserted by Coq. 146 | 147 | #
# 148 | *) 149 | 150 | Lemma leq0n (n : nat) : 0 <= n. 151 | Abort. 152 | 153 | (** 154 | #
# 155 | 156 | Equality statement between tests reads as "if and only if". 157 | 158 | #
# 159 | *) 160 | 161 | Print Nat.sub. 162 | 163 | Lemma eqn_leq (m n : nat) : m == n = (m <= n) && (n <= m). 164 | Abort. 165 | 166 | (** 167 | #
# 168 | 169 | [(_ <= _) && (_ <= _)] has a special notation [(_ <= _ <= _)] 170 | 171 | #
# 172 | *) 173 | 174 | Lemma eqn_leq (m n : nat) : m == n = (m <= n <= m). 175 | Abort. 176 | 177 | (** 178 | #
# 179 | 180 | #


# 181 | 182 | #

(notes)
# 183 | This slide corresponds to 184 | section 2.1 of 185 | #the Mathematical Components book# 186 | #
# 187 | 188 | #


# 189 | #

# 190 | ---------------------------------------------------------- 191 | #
# 192 | ** Proofs by computation 193 | 194 | Our statements are programs. Hence they compute! 195 | 196 | The [by[]] tactic solves trivial goal (mostly) by 197 | computation. 198 | 199 | << 200 | Fixpoint addn n m := 201 | if n is p.+1 then (addn p m).+1 else m. 202 | >> 203 | #
# 204 | *) 205 | 206 | Goal 2 + 2 = 4. 207 | Proof. by []. Qed. 208 | 209 | 210 | Lemma addSn m n : m.+1 + n = (m + n).+1. 211 | Proof. by []. Qed. 212 | 213 | 214 | (** 215 | #
# 216 | << 217 | Fixpoint subn m n : nat := 218 | match m, n with 219 | | p.+1, q.+1 => subn p q 220 | | _ , _ => m 221 | end. 222 | >> 223 | #
# 224 | *) 225 | 226 | Goal 2 - 4 = 0. 227 | Proof. by []. Qed. 228 | 229 | Lemma subn0 m n : m.+1 - n.+1 = m - n. 230 | Proof. by []. Qed. 231 | 232 | (** 233 | #
# 234 | << 235 | Definition leq m n := m - n == 0. 236 | >> 237 | #
# 238 | *) 239 | 240 | Goal 0 <= 4. 241 | Proof. by []. Qed. 242 | 243 | (** 244 | #
# 245 | 246 | [_ < _] is just a notation for [_.+1 <= _]. 247 | 248 | #
# 249 | *) 250 | 251 | Goal 3 < 3 = false. 252 | Proof. by []. Qed. 253 | 254 | Goal 4 <= 3 = false. 255 | Proof. by []. Qed. 256 | 257 | Lemma leq0n n : 0 <= n. 258 | Proof. by []. Qed. 259 | 260 | Lemma ltn0 n : n.+1 <= 0 = false. 261 | Proof. by []. Qed. 262 | 263 | Lemma ltnS m n : (m.+1 <= n.+1) = (m <= n). 264 | Proof. by []. Qed. 265 | 266 | (** 267 | #
# 268 | 269 | Notice the naming convention. 270 | 271 | #
# 272 | *) 273 | 274 | Print negb. 275 | Locate "~~". 276 | Search negb inside ssr.ssrbool. 277 | 278 | Lemma negbK (b : bool) : ~~ (~~ b) = b. 279 | Proof. Fail by []. Abort. 280 | 281 | (** 282 | #
# 283 | 284 | It is not always the case the computation solves all our 285 | problems. In particular here there are no constructors to 286 | consume, hence computation is stuck. 287 | 288 | To prove [negbK] we need a case split. 289 | 290 | #


# 291 | 292 | 293 | #

(notes)
# 294 | This slide corresponds to 295 | section 2.2.1 of 296 | #the Mathematical Components book# 297 | #
# 298 | 299 | #


# 300 | #

# 301 | 302 | 303 | ---------------------------------------------------------- 304 | #
# 305 | ** Proofs by case analysis 306 | 307 | The proof of [negbK] requires a case analysis: given that 308 | [b] is of type bool, it can only be [true] or [false]. 309 | 310 | The [case: term] command performs this proof step. 311 | 312 | #
# 313 | *) 314 | 315 | Lemma negbK b : ~~ (~~ b) = b. 316 | Proof. 317 | case: b. 318 | by []. 319 | by []. 320 | Qed. 321 | 322 | Lemma andbC (b1 b2 : bool) : b1 && b2 = b2 && b1. 323 | Proof. 324 | by case: b1; case: b2. 325 | Qed. 326 | 327 | Lemma orbN b : b || ~~ b. 328 | Proof. 329 | by case: b. 330 | Qed. 331 | 332 | (** 333 | #
# 334 | 335 | The constructors of [bool] have no arguments, but for 336 | example the second constructor of [nat] has one. 337 | 338 | In this case one has to complement the command by supplying 339 | names for these arguments. 340 | 341 | #
# 342 | *) 343 | 344 | Lemma leqn0 n : (n <= 0) = (n == 0). 345 | Proof. 346 | case: n => [| p]. 347 | by []. 348 | by []. 349 | Qed. 350 | 351 | (** 352 | #
# 353 | 354 | Sometimes case analysis is not enough. 355 | 356 | [[ 357 | Fixpoint muln (m n : nat) : nat := 358 | if m is p.+1 then n + muln p n else 0. 359 | ]] 360 | 361 | #
# 362 | *) 363 | Lemma muln_eq0 m n : 364 | (m * n == 0) = (m == 0) || (n == 0). 365 | Proof. 366 | case: m => [|p]. 367 | by []. 368 | case: n => [|k]; last first. (* rotates the goals *) 369 | by []. 370 | Abort. 371 | 372 | (** 373 | #
# 374 | 375 | We don't know how to prove this yet. 376 | But maybe someone proved it already... 377 | 378 | #
# 379 | *) 380 | Search _ (_ * 0) inside ssrnat. (* :-( *) 381 | Search _ muln 0 inside ssrnat. 382 | Print right_zero. 383 | Search right_zero. 384 | 385 | (** 386 | #
# 387 | 388 | The [Search] command is quite primitive but also 389 | your best friend. 390 | 391 | It takes a head pattern, the whildcard [_] 392 | in the examples above, followed by term or patterns, 393 | and finally a location, in this case the [ssrnat] library. 394 | 395 | Our first attempt was unsuccessful because 396 | standard properies (associativity, communtativity, ...) 397 | are expressed in Mathematical Components using 398 | higher order predicates. 399 | In this way these lemmas are very consistent, but also 400 | harder to find if one does not know that. 401 | 402 | The second hit is what we need to complete the proof. 403 | 404 | #
(notes)
# 405 | This slide corresponds to 406 | sections 2.2.2 and 2.5 of 407 | #the Mathematical Components book# 408 | #
# 409 | 410 | #


# 411 | #

# 412 | 413 | ---------------------------------------------------------- 414 | #
# 415 | ** Proofs by rewriting 416 | 417 | The [rewrite] tactic uses an equation. If offers many 418 | more flags than the ones we will see (hint: check the 419 | Coq user manual, SSReflect chapter). 420 | 421 | #
# 422 | *) 423 | 424 | Lemma muln_eq0 m n : 425 | (m * n == 0) = (m == 0) || (n == 0). 426 | Proof. 427 | case: m => [|p]. 428 | by []. 429 | case: n => [|q]. 430 | rewrite muln0. 431 | by []. 432 | by []. 433 | Qed. 434 | 435 | (** 436 | #
# 437 | 438 | Let's now look at another example to learn more 439 | about [rewrite]. 440 | 441 | #
# 442 | *) 443 | Lemma leq_mul2l m n1 n2 : 444 | (m * n1 <= m * n2) = (m == 0) || (n1 <= n2). 445 | Proof. 446 | rewrite /leq. 447 | (* Search _ muln subn in ssrnat. *) 448 | rewrite -mulnBr. 449 | rewrite muln_eq0. 450 | by []. 451 | Qed. 452 | 453 | (** 454 | #
# 455 | 456 | 457 | #
(notes)
# 458 | This slide corresponds to 459 | section 2.2.3 of 460 | #the Mathematical Components book# 461 | #
# 462 | 463 | #


# 464 | #

# 465 | 466 | ---------------------------------------------------------- 467 | #
# 468 | ** Lesson 2: sum up 469 | 470 | - [by []] trivial proofs (including computation) 471 | - [case: m] case split 472 | - [rewrite t1 -t2 /def] rewrite 473 | 474 | #
# 475 | 476 | *) 477 | -------------------------------------------------------------------------------- /AnIntroductionToSmallScaleReflectionInCoq/section4.v: -------------------------------------------------------------------------------- 1 | (******************************************************************************) 2 | (* Solutions of exercises : Small scale reflection, first examples *) 3 | (******************************************************************************) 4 | 5 | From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat div seq. 6 | From mathcomp Require Import path choice fintype tuple finset. 7 | 8 | Set Implicit Arguments. 9 | Unset Strict Implicit. 10 | Import Prenex Implicits. 11 | 12 | 13 | (******************************************************************************) 14 | (* Exercise 4.1.1 *) 15 | (******************************************************************************) 16 | 17 | (* The proof goes by case analysis on the truth table, applying the *) 18 | (*appropriate constructor of the reflect predicate in every meaningful *) 19 | (*case. The other case ask a proof of False in a conctext where False *) 20 | (*can be obtained as a hypothesis after destructing a conjunction (the *) 21 | (*last case tactic) *) 22 | 23 | Lemma tuto_andP : forall b1 b2 : bool, reflect (b1 /\ b2) (b1 && b2). 24 | Proof. by case; case; constructor; auto; case. Qed. 25 | 26 | (* Again this lemma is a macro for a cacse analysis on the truth table *) 27 | (*of b1 b2 and (b1 || b2) *) 28 | 29 | Lemma tuto_orP : forall b1 b2 : bool, reflect (b1 \/ b2) (b1 || b2). 30 | Proof. by case; case; constructor; auto; case. Qed. 31 | 32 | (******************************************************************************) 33 | (* Exercise 4.1.2 *) 34 | (******************************************************************************) 35 | 36 | (* The first solution follows the hint given in the tutorial. In fact *) 37 | (* it is sufficient to perform case analysis on the constructor of the *) 38 | (*reflect hypothesis as done in the second solution. *) 39 | Lemma tuto_iffP : forall (P Q : Prop) (b : bool), 40 | reflect P b -> (P -> Q) -> (Q -> P) -> reflect Q b. 41 | Proof. by move=> P Q; case; case; constructor; auto. Qed. 42 | 43 | 44 | Lemma alternate_tuto_iffP : forall (P Q : Prop) (b : bool), 45 | reflect P b -> (P -> Q) -> (Q -> P) -> reflect Q b. 46 | Proof. by move=> P Q b; case; constructor; auto. Qed. 47 | 48 | (******************************************************************************) 49 | (* Exercise 4.2.1 *) 50 | (******************************************************************************) 51 | 52 | (*We use a section to factorize the type of elements in the *) 53 | (*sequences*) 54 | 55 | Section Exo_4_2_1. 56 | 57 | Variable A : Type. 58 | 59 | (* We use implicit type annotations to avoid casting quantified *) 60 | (*variables : *) 61 | Implicit Types s : seq A. 62 | Implicit Types x : A. 63 | 64 | (* The type of x has been declared implicit, hence the type of s1 and *) 65 | (*s2 and inferred: *) 66 | Fixpoint tuto_cat s1 s2 := 67 | match s1 with 68 | |[::] => s2 69 | | x :: s' => x :: (tuto_cat s' s2) 70 | end. 71 | 72 | (* Using the ssreflect pattern conditional this code can be shrinked *) 73 | (*into the actual program present in the seq library:*) 74 | Fixpoint alternate_tuto s1 s2 := 75 | if s1 is x :: s' then x :: (tuto_cat s' s2) else s2. 76 | 77 | 78 | (* The proof goes by induction on the first list. We do not need to be *) 79 | (*general with respect to the second list, hence we introduce it *) 80 | (*before the induction. The first case is solved on the fly by the //= *) 81 | (*simple+solve switch, and x and s1 are introduced in the remaining *) 82 | (*goal. The last hypothesis is directly rewritten without being *) 83 | (*introduced, generating an equality between convertible terms. This *) 84 | (*equality being trivial it is solved by the prenex 'by' tactic *) 85 | Lemma tuto_size_cat : forall s1 s2, 86 | size (s1 ++ s2) = size s1 + size s2. 87 | Proof. by move=> s1 s2; elim: s1 => //= x s1 ->. Qed. 88 | 89 | (* We use again a pattern conditional *) 90 | Fixpoint tuto_last (A : Type)(x : A)(s : seq A) {struct s} := 91 | if s is x' :: s' then tuto_last x' s' else x. 92 | 93 | (* Again an induction on the first list *) 94 | Lemma tuto_last_cat : forall x s1 s2, 95 | last x (s1 ++ s2) = last (last x s1) s2. 96 | Proof. by move=> x s1 s2; elim: s1 x => [|y s1 IHs] x //=; rewrite IHs. Qed. 97 | 98 | Fixpoint tuto_take n s {struct s} := 99 | match s, n with 100 | | x :: s', n'.+1 => x :: take n' s' 101 | | _, _ => [::] 102 | end. 103 | 104 | 105 | (* Here we have two options: a recursion on the nat or on the seq.*) 106 | (* Decreazing on the seq has better compositional properties: it *) 107 | (*allows more fixpoints further defined to decrease structurally *) 108 | Fixpoint tuto_drop n s {struct s} := 109 | match s, n with 110 | | _ :: s', n'.+1 => drop n' s' 111 | | _, _ => s 112 | end. 113 | 114 | Definition tuto_rot n s := drop n s ++ take n s. 115 | 116 | Lemma tuto_rot_addn : forall m n s, m + n <= size s -> 117 | rot (m + n) s = rot m (rot n s). 118 | Proof. 119 | (* We first transform the inequality hypothesis into a case *) 120 | (*disjunction between equality or strict inequality. *) 121 | move=> m n s; rewrite leq_eqVlt. 122 | (* Then a view allows to proceed to the case analysis by generating *) 123 | (* two goals *) 124 | case/predU1P=> [Emn|Hmn]. 125 | (* 1st case: equality. We rewrite the hypothesis, then the rot_size *) 126 | (*lemma (try to guess the name of the lemma according to the ssr *) 127 | (*discipline, or use the Search command, for instance: 128 | Search _ rot size.*) 129 | rewrite Emn rot_size. 130 | (* Again, use the Search command:*) 131 | Search _ (rot _ (rot _ _) = _). 132 | rewrite rot_add_mod -Emn ?leq_addr ?leq_addl //. 133 | by rewrite Emn leqnn rot_size. 134 | (* a more elegant proof of this last step would be: 135 | by rewrite -{1}(rotrK m s) /rotr -Emn addKn. 136 | It is worth also browsing the source of the library to discover new 137 | functions that could help you, like rotr is this case *) 138 | (* Remember rot is programmed from cat take and drop: 139 | Search _ cat take drop.*) 140 | rewrite -{1}(cat_take_drop n s) /rot !take_cat !drop_cat. 141 | Search _ (size (take _ _) = _). 142 | have Hns : n <= size s by apply: leq_trans (leq_addl _ _) (ltnW Hmn). 143 | rewrite !(size_takel Hns). 144 | (* We directly rewrite the condition proved forward *) 145 | have -> : m + n < n = false by rewrite ltnNge leq_addl. 146 | rewrite size_drop. 147 | have -> : m < size s - n by rewrite ltnNge leq_subLR -ltnNge addnC. 148 | by rewrite addnK catA. 149 | Qed. 150 | 151 | (* Look at the source of the seq.v file for a shorter proof! *) 152 | 153 | End Exo_4_2_1. 154 | (******************************************************************************) 155 | (* Exercise 4.2.2 *) 156 | (******************************************************************************) 157 | 158 | Section Exo_4_2_2. 159 | 160 | Variable A : Type. 161 | 162 | (* We use implicit type annotations to avoid casting quantified *) 163 | (*variables : *) 164 | Implicit Types s : seq A. 165 | Implicit Types x : A. 166 | 167 | (* We fix an arbitrary predicate a *) 168 | Variable a : pred A. 169 | 170 | (* We use again a pattern conditional *) 171 | Fixpoint tuto_count s := if s is x :: s' then a x + tuto_count s' else 0. 172 | 173 | Lemma tuto_count_predUI : forall a1 a2 s, 174 | count (predU a1 a2) s + count (predI a1 a2) s = count a1 s + count a2 s. 175 | Proof. 176 | move=> a1 a2. 177 | (* The proof goes by induction on the list. In the inductive case, we *) 178 | (*introduce the head, the tail, and the induction hypothesis. After 179 | simplification, the case case is trivial, and since simplification is 180 | also usefull in the inductive case, the //= swith both simplifies and 181 | closes the first goal.*) 182 | elim=> [|x s IHs] //=. 183 | (* Now a couple of arithmetic rewriting before using the inductive hypothesis*) 184 | rewrite addnCA -addnA addnA addnC IHs -!addnA. 185 | (* The nat_congr tactic (and Ltac tactic defined in ssrnat) normalizes *) 186 | (*both sides of a nat equality to be able to use some congruence, here *) 187 | (*with count a1 s + _ *) 188 | nat_congr. 189 | (* We use it a second time to eliminate (count a2 s) *) 190 | nat_congr. 191 | (* Now it is only a matter of truth table *) 192 | by case (a1 x); case (a2 x). 193 | Qed. 194 | 195 | (* After closing the section, tuto_count_filter has the type required *) 196 | (* by the exercise *) 197 | Lemma tuto_count_filter : forall s, count a s = size (filter a s). 198 | Proof. by elim=> [|x s IHs] //=; rewrite IHs; case (a x). Qed. 199 | 200 | End Exo_4_2_2. 201 | 202 | (******************************************************************************) 203 | (* Exercise 4.3.1 *) 204 | (******************************************************************************) 205 | 206 | Section Exo_4_3_1. 207 | 208 | (* We fix an arbitrary eqType T *) 209 | Variable T : eqType. 210 | 211 | (* We use implicit type annotations to avoid casting quantified *) 212 | (*variables *) 213 | Implicit Types x y : T. 214 | Implicit Type b : bool. 215 | 216 | Lemma tuto_eqxx : forall x, x == x. 217 | Proof. by move=> x; apply/eqP. Qed. 218 | 219 | (* First solution *) 220 | Lemma tuto_predU1l : forall x y b, x = y -> (x == y) || b. 221 | Proof. by move=> x y b exy; rewrite exy eqxx. Qed. 222 | 223 | (* Second solution. The syntax move/eqP-> is equivalent to *) 224 | (*move/eqP => ->, i.e. it applies the view to the top element of the *) 225 | (*goal stack and rewrites the term obtains immediatly without *) 226 | (*introducing it in the context. Similarily, move-> is equaivalent to *) 227 | (*move=> -> and rewrite the top element of the goal (which should be *) 228 | (*an equality) in the rest of the goal, without introducing it in the context *) 229 | 230 | Lemma tuto_predU1l_alt_proof : forall x y b, x = y -> (x == y) || b. 231 | Proof. by move=> x y b; move/eqP->. Qed. 232 | 233 | (* The proof script is the same in both branches thanks to the *) 234 | (* symmetry of the view mechanism. Remeber the 'by' tactical contains *) 235 | (* the 'split' tactic and hence solves the goal after the last *) 236 | (* assumption interpretation move/eqP. *) 237 | Lemma tuto_predD1P : forall x y b, reflect (x <> y /\ b) ((x != y) && b). 238 | Proof. 239 | by move=> x y b; apply: (iffP andP) => [] []; move/eqP. 240 | Qed. 241 | 242 | (* Remember that _ != _ denotes the boolean disequality *) 243 | (* Coq's unification is able to infer from the goal the arguments to *) 244 | (* provide to the eqP lemma*) 245 | Lemma tuto_eqVneq : forall x y, {x = y} + {x != y}. 246 | Proof. by move=> x y; case: eqP; [left | right]. Qed. 247 | 248 | End Exo_4_3_1. -------------------------------------------------------------------------------- /AnSsreflectTutorial/tutorial_Gonthier_LeRoux.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import ssreflect ssrbool ssrfun eqtype ssrnat div. 2 | 3 | (*****************************************************************************) 4 | (* *) 5 | (* This is the proof script associated with the tutorial *) 6 | (* *) 7 | (* An Ssreflect Tutorial *) 8 | (* Georges Gonthier Stéphane Le Roux *) 9 | (* *) 10 | (* https://hal.inria.fr/inria-00407778 *) 11 | (* *) 12 | (*****************************************************************************) 13 | 14 | 15 | (*****************************************************************************) 16 | (* *) 17 | (* Hilbert's axiom S *) 18 | (* *) 19 | (*****************************************************************************) 20 | 21 | Section HilbertSaxiom. 22 | 23 | Variables A B C : Prop. 24 | 25 | Lemma HilbertS : (A -> B -> C) -> (A -> B) -> A -> C. 26 | Proof. 27 | move=> hAiBiC hAiB hA. 28 | move: hAiBiC. 29 | apply. 30 | by []. 31 | by apply: hAiB. 32 | Qed. 33 | 34 | Hypotheses (hAiBiC : A -> B -> C) (hAiB : A -> B) (hA : A). 35 | 36 | Lemma HilbertS2 : C. 37 | Proof. 38 | apply: hAiBiC; first by apply: hA. 39 | exact: hAiB. 40 | Qed. 41 | 42 | Lemma HilbertS3 : C. 43 | Proof. by apply: hAiBiC; last exact: hAiB. Qed. 44 | 45 | Check (hAiB hA). 46 | 47 | Lemma HilbertS4 : C. 48 | Proof. exact: (hAiBiC _ (hAiB _)). Qed. 49 | 50 | Lemma HilbertS5 : C. 51 | Proof. exact: hAiBiC (hAiB _). Qed. 52 | 53 | Lemma HilbertS6 : C. 54 | Proof. exact: HilbertS5. Qed. 55 | 56 | Print HilbertS5. 57 | 58 | Print HilbertS2. 59 | 60 | Print HilbertS. 61 | 62 | Check HilbertS. 63 | 64 | End HilbertSaxiom. 65 | 66 | Print HilbertS5. 67 | 68 | 69 | (*****************************************************************************) 70 | (* *) 71 | (* Logical connectives *) 72 | (* *) 73 | (*****************************************************************************) 74 | 75 | Print bool. 76 | 77 | Section Symmetric_Conjunction_Disjunction. 78 | 79 | Lemma andb_sym : forall A B : bool, A && B -> B && A. 80 | Proof. 81 | case. 82 | by case. 83 | by []. 84 | Qed. 85 | 86 | Lemma andb_sym2 : forall A B : bool, A && B -> B && A. 87 | Proof. by case; case. Qed. 88 | 89 | Lemma andb_sym3 : forall A B : bool, A && B -> B && A. 90 | Proof. by do 2! case. Qed. 91 | 92 | Variables (C D : Prop) (hC : C) (hD : D). 93 | 94 | Check (and C D). 95 | 96 | Print and. 97 | 98 | Check conj. 99 | 100 | Check (conj hC hD). 101 | 102 | Lemma and_sym : forall A B : Prop, A /\ B -> B /\ A. 103 | Proof. by move=> A1 B []. Qed. 104 | 105 | Check or. 106 | 107 | Check or_introl. 108 | 109 | Lemma or_sym : forall A B : Prop, A \/ B -> B \/ A. 110 | Proof. by move=> A B [hA | hB]; [apply: or_intror | apply: or_introl]. Qed. 111 | 112 | Lemma or_sym2 : forall A B : bool, A \/ B -> B \/ A. 113 | Proof. by move=> [] [] AorB; apply/orP; move/orP : AorB. Qed. 114 | 115 | End Symmetric_Conjunction_Disjunction. 116 | 117 | Section R_sym_trans. 118 | 119 | Variables (D : Type) (R : D -> D -> Prop). 120 | 121 | Hypothesis R_sym : forall x y, R x y -> R y x. 122 | 123 | Hypothesis R_trans : forall x y z, R x y -> R y z -> R x z. 124 | 125 | Lemma refl_if : forall x : D, (exists y, R x y) -> R x x. 126 | Proof. 127 | move=> x [y Rxy]. 128 | by apply: R_trans _ (R_sym _ y _). 129 | Qed. 130 | 131 | End R_sym_trans. 132 | 133 | Section Smullyan_drinker. 134 | 135 | Variables (D : Type)(P : D -> Prop). 136 | 137 | Hypothesis (d : D) (EM : forall A, A \/ ~A). 138 | 139 | Lemma drinker : exists x, P x -> forall y, P y. 140 | Proof. 141 | case: (EM (exists y, ~P y)) => [[y notPy]| nonotPy]; first by exists y. 142 | exists d => _ y; case: (EM (P y)) => // notPy. 143 | by case: nonotPy; exists y. 144 | Qed. 145 | 146 | End Smullyan_drinker. 147 | 148 | Section Equality. 149 | 150 | Variable f : nat -> nat. 151 | Hypothesis f00 : f 0 = 0. 152 | 153 | Lemma fkk : forall k, k = 0 -> f k = k. 154 | Proof. 155 | move=> k k0. 156 | by rewrite k0. 157 | Qed. 158 | 159 | Lemma fkk2 : forall k, k = 0 -> f k = k. 160 | Proof. by move=> k ->. Qed. 161 | 162 | Variable f10 : f 1 = f 0. 163 | 164 | Lemma ff10 : f (f 1) = 0. 165 | Proof. by rewrite f10 f00. Qed. 166 | 167 | Variables (D : eqType) (x y : D). 168 | 169 | Lemma eq_prop_bool : x = y -> x == y. 170 | Proof. by move/eqP. Qed. 171 | 172 | Lemma eq_bool_prop : x == y -> x = y. 173 | Proof. by move/eqP. Qed. 174 | 175 | End Equality. 176 | 177 | Section Using_Definition. 178 | 179 | Variable U : Type. 180 | 181 | Definition set := U -> Prop. 182 | 183 | Definition subset (A B : set) := forall x, A x -> B x. 184 | 185 | Definition transitive (T : Type) (R : T -> T -> Prop) := 186 | forall x y z, R x y -> R y z -> R x z. 187 | 188 | Lemma subset_trans : transitive set subset. 189 | Proof. 190 | rewrite /transitive /subset => x y z subxy subyz t xt. 191 | by apply: subyz; apply: subxy. 192 | Qed. 193 | 194 | Lemma subset_trans2 : transitive set subset. 195 | Proof. 196 | move=> x y z subxy subyz t. 197 | by move/subxy; move/subyz. 198 | Qed. 199 | 200 | End Using_Definition. 201 | 202 | (*****************************************************************************) 203 | (* *) 204 | (* Arithmetic for Euclidean division *) 205 | (* *) 206 | (*****************************************************************************) 207 | 208 | Section Arithmetic. 209 | 210 | Check nat. 211 | 212 | Print nat. 213 | 214 | Lemma three : S (S (S O)) = 3 /\ 2 = 0.+1.+1. 215 | Proof. by []. Qed. 216 | 217 | Print plus. 218 | 219 | Lemma concrete_plus : plus 16 64 = 80. 220 | Proof. (* simpl. *) by []. Qed. 221 | 222 | Lemma concrete_plus_bis : 16 + 64 = 80. 223 | Proof. (* simpl. *) by []. Qed. 224 | 225 | Print mult. 226 | 227 | Print le. 228 | 229 | Lemma concrete_le : le 1 3. 230 | Proof. 231 | by apply: (PeanoNat.Nat.le_trans _ 2); apply: PeanoNat.Nat.le_succ_diag_r. 232 | Qed. 233 | 234 | Lemma concrete_big_le : le 16 64. 235 | Proof. by auto 50. Qed. 236 | 237 | Print leq. 238 | 239 | Print subn. 240 | 241 | Print subn_rec. 242 | 243 | Lemma concrete_big_leq : 0 <= 51. 244 | Proof. by []. Qed. 245 | 246 | Lemma semi_concrete_leq : forall n m, n <= m -> 51 + n <= 51 + m. 247 | Proof. by []. Qed. 248 | 249 | Lemma concrete_arith : (50 < 100) && (3 + 4 < 3 * 4 <= 17 - 2). 250 | Proof. by []. Qed. 251 | 252 | Print nat_ind. 253 | 254 | Lemma plus_commute : forall n1 m1, n1 + m1 = m1 + n1. 255 | Proof. 256 | elim=> [m1 | n1 IHn1 m1]. 257 | by elim: m1 => // m1 IHm1; rewrite -[0 + m1.+1]/(0 + m1).+1 IHm1. 258 | rewrite -[n1.+1 + m1]/(n1 + m1).+1 IHn1. 259 | by elim: m1 => // m1 IHm1; rewrite -[m1.+1 + n1]/(m1 + n1).+1 IHm1. 260 | Qed. 261 | 262 | Check edivn_rec. 263 | 264 | Print edivn_rec. 265 | 266 | Print edivn. 267 | 268 | Print edivn_spec. 269 | 270 | Lemma edivnP : forall m d, edivn_spec m d (edivn m d). 271 | Proof. 272 | rewrite /edivn => m [|d] //=; rewrite -{1}[m]/(0 * d.+1 + m). 273 | elim: m {-2}m 0 (leqnn m) => [|n IHn] [|m] q //=; rewrite ltnS => le_mn. 274 | rewrite subn_if_gt; case: ltnP => [// | le_dm]. 275 | rewrite -{1}(subnK le_dm) [_ + d]addnC -addSn addnA -mulSnr; apply: IHn. 276 | apply: leq_trans le_mn; exact: leq_subr. 277 | Qed. 278 | 279 | Check nat_ind. 280 | 281 | Check ltnP. 282 | 283 | Print ltn_xor_geq. 284 | 285 | Lemma edivn_eq : forall d q r, r < d -> edivn (q * d + r) d = (q, r). 286 | Proof. 287 | move=> d q r lt_rd; have d_pos: 0 < d by exact: leq_trans lt_rd. 288 | case: edivnP lt_rd => q' r'; rewrite d_pos /=. 289 | wlog: q q' r r' / q <= q' by case (ltnP q q'); last symmetry; eauto. 290 | rewrite leq_eqVlt; case: eqP => [-> _|_] /=; first by move/addnI->. 291 | rewrite -(leq_pmul2r d_pos); move/leq_add=> Hqr Eqr _; move/Hqr {Hqr}. 292 | by rewrite addnS ltnNge mulSn -addnA Eqr addnCA addnA leq_addr. 293 | Qed. 294 | 295 | (*****************************************************************************) 296 | (* *) 297 | (* Parametric type families and alternative specifications *) 298 | (* *) 299 | (*****************************************************************************) 300 | 301 | 302 | Check edivn_spec. 303 | 304 | Print edivn_spec. 305 | 306 | CoInductive edivn_spec_right : nat -> nat -> nat * nat -> Type := 307 | EdivnSpec_right m d q r of m = q * d + r & (d > 0) ==> (r < d) : 308 | edivn_spec_right m d (q, r). 309 | 310 | CoInductive edivn_spec_left (m d : nat)(qr : nat * nat) : Type := 311 | EdivnSpec_left of m = (fst qr) * d + (snd qr) & (d > 0) ==> (snd qr < d) : 312 | edivn_spec_left m d qr. 313 | 314 | 315 | Lemma edivnP_right : forall m d, edivn_spec_right m d (edivn m d). 316 | Admitted. 317 | 318 | Lemma edivnP_left : forall m d, edivn_spec_left m d (edivn m d). 319 | Admitted. 320 | 321 | 322 | Lemma edivn_eq_right : forall d q r, r < d -> edivn (q * d + r) d = (q, r). 323 | Proof. 324 | move=> d q r lt_rd; have d_pos: 0 < d by exact: leq_trans lt_rd. 325 | set m := q * d + r; have: m = q * d + r by []. 326 | set d' := d; have: d' = d by []. 327 | case: (edivnP_right m d') => {d'}m d' q' r' -> lt_r'd' d'd q'd'r'. 328 | move: q'd'r' lt_r'd' lt_rd; rewrite d'd d_pos {d'd m} /=. 329 | wlog: q q' r r' / q <= q' by case (ltnP q q'); last symmetry; eauto. 330 | rewrite leq_eqVlt; case: eqP => [-> _|_] /=; first by move/addnI->. 331 | rewrite -(leq_pmul2r d_pos); move/leq_add=> Hqr Eqr _; move/Hqr {Hqr}. 332 | by rewrite addnS ltnNge mulSn -addnA -Eqr addnCA addnA leq_addr. 333 | Qed. 334 | 335 | Lemma edivn_eq_left : forall d q r, r < d -> edivn (q * d + r) d = (q, r). 336 | Proof. 337 | move=> d q r lt_rd; have d_pos: 0 < d by exact: leq_trans lt_rd. 338 | case: (edivnP_left (q * d + r) d) lt_rd; rewrite d_pos /=. 339 | set q':= (edivn (q * d + r) d).1; set r':= (edivn (q * d + r) d).2. 340 | rewrite (surjective_pairing (edivn (q * d + r) d)) -/q' -/r'. 341 | wlog: q r q' r' / q <= q' by case (ltnP q q'); last symmetry; eauto. 342 | rewrite leq_eqVlt; case: eqP => [-> _|_] /=; first by move/addnI->. 343 | rewrite -(leq_pmul2r d_pos); move/leq_add=> Hqr Eqr _; move/Hqr {Hqr}. 344 | by rewrite addnS ltnNge mulSn -addnA Eqr addnCA addnA leq_addr. 345 | Qed. 346 | 347 | End Arithmetic. 348 | -------------------------------------------------------------------------------- /SummerSchoolSophia/lesson4.v: -------------------------------------------------------------------------------- 1 | 2 | From mathcomp Require Import all_ssreflect. 3 | 4 | Set Implicit Arguments. 5 | Unset Strict Implicit. 6 | Unset Printing Implicit Defensive. 7 | 8 | (** 9 | 10 | ---------------------------------------------------------- 11 | #
# 12 | ** Lesson 4: summary 13 | 14 | - Curry Howard: the big picture 15 | + dependent function space 16 | - Predicates and connectives 17 | + introduction 18 | + elimination 19 | - Induction 20 | - Consistency 21 | - Dependent elimination 22 | 23 | #
# 24 | 25 | ---------------------------------------------------------- 26 | #
# 27 | ** Curry Howard 28 | 29 | We link typed programs to statements with a proof. 30 | 31 | Let's play a game in which we use inductive types 32 | as our satements. 33 | 34 | #
# 35 | *) 36 | 37 | Check nat : Type. 38 | 39 | Definition zero : nat := 0. 40 | 41 | Lemma zero_bis : nat. 42 | Proof. 43 | apply: 0. 44 | Qed. 45 | 46 | Print zero. 47 | Print zero_bis. 48 | 49 | 50 | (** 51 | #
# 52 | 53 | We learn that 0 is a term of type nat, but Coq 54 | also accepts it as a proof of nat. 55 | 56 | #
# 57 | In type theory: [p] is a proof of [T] 58 | means that [p] inhabits the type [T]. 59 | #
# 60 | 61 | Now let's look at the function space. 62 | 63 | #
# 64 | *) 65 | 66 | Check nat -> nat : Type. 67 | 68 | Definition silly : nat -> nat := fun x => x. 69 | 70 | Lemma sillier : nat -> nat. 71 | Proof. move=> x. apply: x. Qed. 72 | 73 | Print silly. 74 | Print sillier. 75 | 76 | (** 77 | #
# 78 | 79 | The function space [->] can represent implication. 80 | An inhabitant of [A -> B] is a function turning 81 | a proof of [A] into a proof of [B] (a program 82 | taking in input a term of type [A] and returning 83 | a term of type [B]). 84 | 85 | The function space of type theory is *dependent*. 86 | 87 | #
# 88 | *) 89 | 90 | Section DependentFunction. 91 | 92 | Variable P : nat -> Type. 93 | Variable p1 : P 1. 94 | 95 | 96 | Check forall x, P x. 97 | Check forall x : nat, P 1. 98 | 99 | Check fun x : nat => p1. 100 | 101 | 102 | (** 103 | #
# 104 | 105 | We managed to build (introduce) an arrow and a forall using [fun]. 106 | Let's see how we can use (eliminate) an arrow or a forall. 107 | 108 | #
# 109 | *) 110 | 111 | Check factorial. 112 | Check factorial 2. 113 | 114 | Variable px1 : forall x, P x.+1. 115 | 116 | Check px1. 117 | Check px1 3. 118 | 119 | End DependentFunction. 120 | 121 | (** 122 | #
# 123 | 124 | Following the Curry Howard correspondence *application* 125 | lets one call a function [f : A -> B] on [a : A] to 126 | obtain a term of type [B]. If the type of [f] is 127 | a dependent arrow (forall) [f : forall x : A, B x] 128 | then the argument [a] appears in the type of 129 | term we obtain, that is [f a] has type [B a]. 130 | 131 | In other words application instantiates universally 132 | quantified lemmas and implements modus ponens. 133 | 134 | Lemmas can be seen as views to transform assumptions. 135 | 136 | #
# 137 | *) 138 | 139 | Section Views. 140 | 141 | Variable P : nat -> Type. 142 | Variable Q : nat -> Type. 143 | Variable p2q : forall x, P x -> Q x. 144 | 145 | Goal P 3 -> True. 146 | Proof. 147 | move=> (*/p2q*) p3. 148 | Abort. 149 | 150 | End Views. 151 | 152 | (** 153 | #
# 154 | 155 | So far we used [nat] (and [P]) as a predicate and [->] for implication. 156 | 157 | Can we use inductive types to model other predicates or connectives? 158 | 159 | #


# 160 | 161 | #

(notes)
# 162 | 163 | This slide corresponds to 164 | section 4.x of 165 | #the Mathematical Components book# 166 | #
# 167 | 168 | 169 | #
# 170 | 171 | ---------------------------------------------------------- 172 | #
# 173 | ** Predicates and connectives 174 | 175 | Let's start with #$$ \top $$# 176 | 177 | Note: here the label [Prop] could be a synonym of [Type]. 178 | 179 | #
# 180 | *) 181 | 182 | Print True. 183 | 184 | Definition trivial1 : True := I. 185 | 186 | Definition trivial2 : True -> nat := 187 | fun t => 188 | match t with I => 3 end. 189 | 190 | Lemma trivial3 : True -> nat. 191 | Proof. 192 | move=> t. case: t. apply: 3. 193 | Qed. 194 | 195 | (** 196 | #
# 197 | 198 | Now let's look at #$$ \bot $$# 199 | 200 | #
# 201 | *) 202 | 203 | Print False. 204 | 205 | Fail Definition hard1 : False := what. 206 | 207 | Definition ex_falso A : False -> A := 208 | fun abs => match abs with end. 209 | 210 | Lemma ex_falso2 A : False -> A. 211 | Proof. 212 | move=> abs. case: abs. 213 | Qed. 214 | 215 | (** 216 | #
# 217 | 218 | Connectives: #$$ \land $$# and #$$\lor $$# 219 | 220 | #
# 221 | *) 222 | 223 | Section Connectives. 224 | 225 | Print and. 226 | 227 | Variable A : Prop. 228 | Variable B : Prop. 229 | Variable C : Prop. 230 | 231 | Variable a : A. 232 | Variable b : B. 233 | 234 | Check conj a b. 235 | 236 | Definition and_elim_left : and A B -> A := 237 | fun ab => match ab with conj a b => a end. 238 | 239 | 240 | Lemma and_elim_left2 : and A B -> A. 241 | Proof. case=> l r. apply: l. Qed. 242 | 243 | Print or. 244 | 245 | Check or_introl a : or A B. 246 | Check or_intror b : or A B. 247 | 248 | Definition or_elim : 249 | A \/ B -> (A -> C) -> (B -> C) -> C := 250 | fun aob a2c b2c => 251 | match aob with 252 | | or_introl a => a2c a 253 | | or_intror b => b2c b 254 | end. 255 | 256 | Lemma or_elim_example : A \/ B -> C. 257 | Proof. 258 | move=> aob. 259 | case: aob. 260 | Abort. 261 | 262 | (** 263 | #
# 264 | 265 | Quantifier #$$ \exists $$# 266 | 267 | #
# 268 | *) 269 | 270 | Print ex. 271 | 272 | Lemma ex_elim P : (exists x : A, P x) -> True. 273 | Proof. 274 | case => x px. 275 | Abort. 276 | 277 | End Connectives. 278 | 279 | (** 280 | #
# 281 | 282 | #


# 283 | 284 | #

(notes)
# 285 | 286 | This slide corresponds to 287 | section 4.x of 288 | #the Mathematical Components book# 289 | #
# 290 | 291 | 292 | #
# 293 | ---------------------------------------------------------- 294 | #
# 295 | ** Induction 296 | 297 | We want to prove theorems by induction, right? 298 | Hence there must be a term that corresponds to the induction principle. 299 | This term is a recursive function. 300 | 301 | Note: [Fixpoint] is just sugar for [Definition] followed by [fix]. 302 | 303 | #
# 304 | *) 305 | 306 | About nat_ind. 307 | 308 | Definition ind : 309 | forall P : nat -> Prop, 310 | P 0 -> (forall n : nat, P n -> P n.+1) -> forall n : nat, P n := 311 | fun P p0 pS => 312 | fix IH n : P n := 313 | match n with 314 | | O => p0 315 | | S p => pS p (IH p) 316 | end. 317 | 318 | (** 319 | #
# 320 | 321 | #


# 322 | 323 | #

(notes)
# 324 | 325 | This slide corresponds to 326 | section 4.x of 327 | #the Mathematical Components book# 328 | #
# 329 | 330 | ---------------------------------------------------------- 331 | #
# 332 | ** Consistency 333 | 334 | We give here the intuition why some terms that are in principle 335 | well typed are rejected by Coq and why Coq is consistent. 336 | 337 | #
# 338 | *) 339 | Print False. 340 | Print True. 341 | (** 342 | #
# 343 | 344 | What does it mean that [t : T] and [T] is not [False]? 345 | 346 | #
# 347 | *) 348 | Check (match 3 with O => I | S _ => I end) : True. 349 | (** 350 | #
# 351 | 352 | Constructors are not the only terms that can inhabit a type. 353 | Hence we cannot simply look at terms, but we could look at 354 | their normal form. 355 | 356 | Subject reduction: [t : T] and [t ~> t1] then [t1 : T]. 357 | We claim there is not such [t1] (normal form) that 358 | inhabits [False]. 359 | 360 | We have to reject [t] that don't have a normal form. 361 | 362 | Exaustiveness of pattern matching: 363 | 364 | #
# 365 | *) 366 | 367 | Lemma helper x : S x = 0 -> False. Proof. by []. Qed. 368 | 369 | Fail Definition partial n : n = 0 -> False := 370 | match n with 371 | | S x => fun p : S x = 0 => helper x p 372 | (* | 0 => fun _ => I*) 373 | end. 374 | 375 | Fail Check partial 0 (erefl 0). (* : False *) 376 | 377 | Fail Compute partial 0 (erefl 0). (* = ??? : False *) 378 | 379 | (** 380 | #
# 381 | 382 | According to Curry Howard this means that in a case 383 | split we did not forget to consider a branch! 384 | 385 | Termination of recursion: 386 | 387 | #
# 388 | *) 389 | 390 | Fail Fixpoint oops (n : nat) : False := oops n.+1. 391 | 392 | Fail Check oops 3. (* : False *) 393 | 394 | Fail Compute oops 3. (* = ??? : False *) 395 | 396 | (** 397 | #
# 398 | 399 | According to Curry Howard this means that we did not 400 | do any circular argument. 401 | 402 | Non termination is subtle since a recursive call could 403 | be hidden in a box. 404 | 405 | #
# 406 | *) 407 | 408 | Fail Inductive hidden := Hide (f : hidden -> False). 409 | 410 | Fail Definition oops (hf : hidden) : False := 411 | match hf with Hide f => f hf end. 412 | 413 | Fail Check oops (Hide oops). (* : False *) 414 | 415 | Fail Compute oops (Hide oops). (* = ??? : False *) 416 | 417 | (** 418 | #
# 419 | 420 | This condition of inductive types is called positivity: 421 | The type of [Hide] would be [(hidden -> False) -> hidden], 422 | where the first occurrence of [hidden] is on the left (negative) 423 | of the arrow. 424 | 425 | #


# 426 | 427 | #

(notes)
# 428 | 429 | This slide corresponds to 430 | section 3.2.3 of 431 | #the Mathematical Components book# 432 | #
# 433 | 434 | 435 | #
# 436 | 437 | ---------------------------------------------------------- 438 | #
# 439 | ** Inductive types with indexes (casse-tête) 440 | 441 | ... and their elimination. 442 | 443 | The intuition, operationally. 444 | 445 | Inductive types can express tricky invariants: 446 | 447 | #
# 448 | *) 449 | 450 | (* Translucent box, we know if it is empty or not without opening it *) 451 | Inductive tbox : bool -> Type := 452 | | Empty : tbox false 453 | | Full (n : nat) : tbox true. 454 | 455 | Check Empty. 456 | Check Full 3. 457 | 458 | Definition default (d : nat) (f : bool) (b : tbox f) : nat := 459 | match b with 460 | | Empty => d 461 | | Full x => x 462 | end. 463 | 464 | (* Why this complication? (believe me, not worth it) *) 465 | Definition get (b : tbox true) : nat := 466 | match b with Full x => x end. 467 | 468 | (* the meat: why is the elimination tricky? *) 469 | Lemma default_usage f (b : tbox f) : 0 <= default 3 b . 470 | Proof. 471 | case: b. 472 | Fail Check @default 3 f Empty. 473 | by []. 474 | by []. 475 | Qed. 476 | 477 | (** 478 | #
# 479 | 480 | Take home: 481 | - the elimination of an inductive data type with indexes 482 | expresses equations between the value of the indexes 483 | in the type of the eliminated term and the value of the 484 | indexes prescribed in the declatation of the inductive data 485 | - the implicit equations are substituted automatically at 486 | elimination time 487 | - working with indexed data is hard, too hard :-/ 488 | - we can still make good use of indexes when we define "spec" lemmas, 489 | argument of the next lecture 490 | 491 | #


# 492 | 493 | #

(notes)
# 494 | 495 | This slide corresponds to 496 | section 4.x of 497 | #the Mathematical Components book# 498 | #
# 499 | 500 | #
# 501 | 502 | ---------------------------------------------------------- 503 | #
# 504 | ** Lesson 4: sum up 505 | 506 | - In Coq terms/types play a double role: 507 | + programs and their types 508 | + statements and their proofs 509 | - Inductives can be used to model predicates and 510 | connectives 511 | - Pattern machind and recursion can model induction 512 | - The empty type is, well, empty, hence Coq is consistent 513 | - Inductives with indexes 514 | 515 | #
# 516 | 517 | *) 518 | -------------------------------------------------------------------------------- /SummerSchoolSophia/lesson1.v: -------------------------------------------------------------------------------- 1 | From mathcomp Require Import all_ssreflect. 2 | 3 | Set Implicit Arguments. 4 | Unset Strict Implicit. 5 | Unset Printing Implicit Defensive. 6 | (** 7 | 8 | ---------------------------------------------------------- 9 | #
# 10 | ** The Coq proof assistant and the Mathematical Components library 11 | 12 | Objective: learn the Coq system in the MC library 13 | 14 | *** Roadmap 15 | 16 | - #lesson 1#: Functions and computations 17 | - #exercise # 18 | 19 | - #lesson 2#: First steps in formal proofs 20 | - #exercise # 21 | 22 | - #lesson 3#: A few more steps in formal proofs 23 | - #exercise # 24 | 25 | - #lesson 4#: Type theory 26 | - #exercise # 27 | 28 | - #lesson 5#: Boolean reflection 29 | - #exercise # 30 | 31 | - #lesson 6#: Real proofs, finally! 32 | - #exercise # 33 | 34 | - #lesson 7#: Generic theories 35 | - #exercise # 36 | 37 | - #lesson 8#: Subtypes 38 | - #exercise # 39 | 40 | *** Teaching material 41 | 42 | - Slides and exercises 43 | #https://www-sop.inria.fr/teams/marelle/coq-18/# 44 | - Coq (#software# 45 | and #user manual#, in particular the chapter about #SSReflect#) 46 | - Mathematical Components 47 | (#software# and 48 | #book#) 49 | 50 | 51 | #


# 52 | 53 | #

(notes)
# 54 | You don't need to install Coq in order to follow this 55 | class, you just need a recent browser thanks to 56 | #jsCoq#. 57 | #
# 58 | 59 | #
# 60 | 61 | ---------------------------------------------------------- 62 | #
# 63 | ** Lesson 1: summary 64 | 65 | - functions 66 | - simple data 67 | - containers 68 | - symbolic computations 69 | - higher order functions and mathematical notations 70 | 71 | #
# 72 | 73 | ---------------------------------------------------------- 74 | #
# 75 | ** Functions 76 | 77 | Functions are built using the [fun .. => ..] syntax. 78 | The command [Check] verifies that a term is well typed. 79 | 80 | #
# 81 | *) 82 | Check (fun n => 1 + n + 1). 83 | (** 84 | #
# 85 | 86 | Notice that the type of [n] was inferred and that 87 | the whole term has type [nat -> nat], where [->] 88 | is the function space. 89 | 90 | Function application is written by writing the function 91 | on the left of the argument (eg, not as in the mathematical 92 | practice). 93 | 94 | #
# 95 | *) 96 | Check 2. 97 | Check (fun n => 1 + n + 1) 2. 98 | (** 99 | #
# 100 | 101 | Notice how [2] has a type that fits, and hence 102 | the type of the function applied to [2] is [nat]. 103 | 104 | Terms (hence functions) can be given a name using 105 | the [Definition] command. The command offers some 106 | syntactic sugar for binding the function arguments. 107 | 108 | #
# 109 | *) 110 | Definition f := (fun n => 1 + n + 1). 111 | (* Definition f n := 1 + n + 1. *) 112 | (* Definition f (n : nat) := 1 + n + 1. *) 113 | (** 114 | #
# 115 | 116 | Named terms can be printed. 117 | 118 | #
# 119 | *) 120 | Print f. 121 | (** 122 | #
# 123 | 124 | Coq is able to compute with terms, in particular 125 | one can obtain the normal form via the [Eval lazy in] 126 | command. 127 | 128 | #
# 129 | *) 130 | Eval lazy in f 2. 131 | (** 132 | #
# 133 | 134 | Notice that "computation" is made of many steps. 135 | In particular [f] has to be unfolded (delta step) 136 | and then the variable substituted for the argument 137 | (beta). 138 | 139 | #
# 140 | *) 141 | Eval lazy delta [f] in f 2. 142 | Eval lazy delta [f] beta in f 2. 143 | (** 144 | #
# 145 | 146 | Nothing but functions (and their types) are built-in in Coq. 147 | All the rest is defined, even [1], [2] and [+] are not primitive. 148 | 149 | #


# 150 | 151 | #

(notes)
# 152 | 153 | This slide corresponds to 154 | section 1.1 of 155 | #the Mathematical Components book# 156 | #
# 157 | 158 | #
# 159 | 160 | ---------------------------------------------------------- 161 | #
# 162 | ** Data types 163 | 164 | Data types can be declared using the [Inductive] command. 165 | 166 | Many of them are already available in the Coq library called 167 | [Prelude] that is automatically loaded. We hence just print 168 | them. 169 | 170 | [Inductive bool := true | false.] 171 | 172 | #
# 173 | *) 174 | Print bool. 175 | (** 176 | #
# 177 | 178 | This command declares a new type [bool] and declares 179 | how the terms (in normal form) of this type are built. 180 | Only [true] and [false] are canonical inhabitants of 181 | [bool]. 182 | 183 | To use a boolean value Coq provides the [if..then..else..] 184 | syntax. 185 | 186 | #
# 187 | *) 188 | Definition twoVtree (b : bool) := if b then 2 else 3. 189 | Eval lazy in twoVtree true. 190 | Eval lazy delta in twoVtree true. 191 | Eval lazy delta beta in twoVtree true. 192 | Eval lazy delta beta iota in twoVtree true. 193 | (** 194 | #
# 195 | 196 | We define a few boolean operators that will come in handy 197 | later on. 198 | 199 | #
# 200 | *) 201 | Definition andb (b1 b2 : bool) := if b1 then b2 else false. 202 | Definition orb (b1 b2 : bool) := if b1 then true else b2. 203 | 204 | Infix "&&" := andb. 205 | Infix "||" := orb. 206 | 207 | Check true && false || false. 208 | (** 209 | #
# 210 | 211 | The [Infix] command lets one declare infix notations. 212 | Precedence and associativity is already declared in the 213 | prelude of Coq, here we just associate the constants 214 | [andb] and [orb] to these notataions. 215 | 216 | Natural numbers are defined similarly to booleans: 217 | 218 | [Inductive nat := O | S (n : nat).] 219 | 220 | #
# 221 | *) 222 | Print nat. 223 | (** 224 | #
# 225 | 226 | Coq provides a special notation for literals, eg [3], 227 | that is just sugar for [S (S (S O))]. 228 | 229 | The Mathematical Components library adds on top of that 230 | the postfix [.+1], [.+2], .. for iterated applications 231 | of [S] to terms other than [O]. 232 | 233 | #
# 234 | *) 235 | Check 3. 236 | Check (fun x => (x + x).+2). 237 | Eval lazy in (fun x => (x + x).+2) 1. 238 | (** 239 | #
# 240 | 241 | In order to use natural numbers Coq provides two 242 | tools. An extended [if..then..else..] syntax to 243 | extract the argument of [S] and the [Fixpoint] 244 | command to define recusrsive functions. 245 | 246 | #
# 247 | *) 248 | Definition pred (n : nat) := 249 | if n is p.+1 then p else 0. 250 | 251 | Eval lazy in pred 7. 252 | (** 253 | #
# 254 | 255 | Notice that [p] is a binder. When the [if..then..else..] 256 | is evaluated, and [n] put in normal form, then if it 257 | is [S t] the variable [p] takes [t] and the then-branch 258 | is taken. 259 | 260 | Now lets define addition using recursion 261 | 262 | #
# 263 | *) 264 | Fixpoint addn n m := 265 | if n is p.+1 then (addn p m).+1 else m. 266 | Infix "+" := addn. 267 | Eval lazy in 3 + 2. 268 | (** 269 | #
# 270 | 271 | The [if..then..else..] syntax is just sugar for 272 | [match..with..end]. 273 | 274 | #
# 275 | *) 276 | Print addn. 277 | (** 278 | #
# 279 | 280 | Let's now write the equality test for natural numbers 281 | 282 | #
# 283 | *) 284 | Fixpoint eqn n m := 285 | match n, m with 286 | | 0, 0 => true 287 | | p.+1, q.+1 => eqn p q 288 | | _, _ => false 289 | end. 290 | Infix "==" := eqn. 291 | Eval lazy in 3 == 4. 292 | (** 293 | #
# 294 | 295 | Other examples are subtraction and order 296 | 297 | #
# 298 | *) 299 | Fixpoint subn m n : nat := 300 | match m, n with 301 | | p.+1, q.+1 => subn p q 302 | | _ , _ => m 303 | end. 304 | 305 | Infix "-" := subn. 306 | 307 | Eval lazy in 3 - 2. 308 | Eval lazy in 2 - 3. (* truncated *) 309 | 310 | Definition leq m n := m - n == 0. 311 | 312 | Infix "<=" := leq. 313 | 314 | Eval lazy in 4 <= 5. 315 | (** 316 | #
# 317 | 318 | #


# 319 | 320 | #

(notes)
# 321 | 322 | All the constants defined in this slide are already 323 | defined in Coq's prelude or in Mathematical Components. 324 | The main difference is that [==] is not specific to 325 | [nat] but overloaded (it works for most data types). 326 | This topic is to be developed in lesson 4. 327 | 328 | This slide corresponds to 329 | section 1.2 of 330 | #the Mathematical Components book# 331 | #
# 332 | 333 | #
# 334 | 335 | #
# 336 | 337 | ---------------------------------------------------------- 338 | #
# 339 | ** Containers 340 | 341 | Containers let one aggregate data, for example to form a 342 | pair or a list. The interesting characteristic of containers 343 | is that they are polymorphic: the same container can be used 344 | to hold terms of many types. 345 | 346 | [Inductive seq (A : Type) := nil | cons (hd : A) (tl : seq A).] 347 | 348 | #
# 349 | *) 350 | Check nil. 351 | Check cons 3 [::]. 352 | (** 353 | #
# 354 | 355 | We learn that [[::]] is a notation for the empty sequence 356 | and that the type parameter [?A] is implicit. 357 | 358 | #
# 359 | *) 360 | Check 1 :: nil. 361 | Check [:: 3; 4; 5 ]. 362 | (** 363 | #
# 364 | 365 | The infix [::] notation stands for [cons]. This one is mostly 366 | used to pattern match a sequence. 367 | 368 | The notation [[:: .. ; .. ]] can be used to form sequences 369 | by separating the elements with [;]. When there are no elements 370 | what is left is [[::]] that is the empty seqeunce. 371 | 372 | And of course we can use sequences with other data types 373 | 374 | #
# 375 | *) 376 | Check [:: 3; 4; 5 ]. 377 | Check [:: true; false; true ]. 378 | (** 379 | #
# 380 | 381 | Let's now define the [size] function. 382 | 383 | #
# 384 | *) 385 | Fixpoint size A (s : seq A) := 386 | if s is _ :: tl then (size tl).+1 else 0. 387 | 388 | Eval lazy in size [:: 1; 8; 34]. 389 | (** 390 | #
# 391 | 392 | Given that the contents of containers are of an 393 | arbitrary type many common operations are parametrized 394 | by functions that are specific to the type of the 395 | contents. 396 | 397 | [[ 398 | Fixpoint map A B (f : A -> B) s := 399 | if s is e :: tl then f e :: map f tl else nil. 400 | ]] 401 | 402 | #
# 403 | *) 404 | Definition l := [:: 1; 2; 3]. 405 | Eval lazy in [seq x.+1 | x <- l]. 406 | (** 407 | #
# 408 | 409 | The #seq# 410 | library of Mathematical Components contains many combinators. Their syntax 411 | is documented in the header of the file. 412 | 413 | #


# 414 | 415 | #

(notes)
# 416 | 417 | This slide corresponds to 418 | section 1.3 of 419 | #the Mathematical Components book# 420 | #
# 421 | 422 | 423 | #
# 424 | 425 | ---------------------------------------------------------- 426 | #
# 427 | ** Symbols 428 | 429 | The section mecanism is used to describe a context under 430 | which definitions are made. Coq lets us not only define 431 | terms, but also compute with them in that context. 432 | 433 | We use this mecanism to talk about symbolic computation. 434 | 435 | #
# 436 | *) 437 | Section symbols. 438 | Variables v : nat. 439 | 440 | Eval lazy in pred v.+1 . 441 | Eval lazy in pred v . 442 | (** 443 | #
# 444 | 445 | Computation can take place in presence of variables 446 | as long as constructors can be consumed. When no 447 | more constructors are available computation is 448 | stuck. 449 | 450 | Let's not look at a very common higher order 451 | function. 452 | 453 | #
# 454 | *) 455 | 456 | Fixpoint foldr A T f (a : A) (s : seq T) := 457 | if s is x :: xs then f x (foldr f a xs) else a. 458 | (** 459 | #
# 460 | 461 | The best way to understand what [foldr] does 462 | is to postulate a variable [f] and compute. 463 | 464 | #
# 465 | *) 466 | 467 | Variable f : nat -> nat -> nat. 468 | 469 | Eval lazy in foldr f 3 [:: 1; 2 ]. 470 | 471 | (** 472 | #
# 473 | 474 | If we plug [addn] in place of [f] we 475 | obtain a term that evaluates to a number. 476 | 477 | #
# 478 | *) 479 | 480 | Eval lazy in foldr addn 3 [:: 1; 2 ]. 481 | 482 | End symbols. 483 | 484 | (** 485 | #
# 486 | 487 | #
(notes)
# 488 | 489 | This slide corresponds to 490 | sections 1.4 and 1.5 of 491 | #the Mathematical Components book# 492 | #
# 493 | 494 | #
# 495 | 496 | ---------------------------------------------------------- 497 | #
# 498 | ** Higher order functions and mathematical notations 499 | 500 | Let's try to write this formula in Coq 501 | 502 | #$$ \sum_{i=1}^n (i * 2 - 1) = n ^ 2 $$# 503 | 504 | We need a bit of infrastruture 505 | 506 | #
# 507 | *) 508 | Fixpoint iota m n := if n is u.+1 then m :: iota m.+1 u else [::]. 509 | 510 | Eval lazy in iota 0 5. 511 | 512 | (** 513 | #
# 514 | 515 | Combining [iota] and [foldr] we can get pretty 516 | close to the LaTeX source for the formula above. 517 | 518 | #
# 519 | *) 520 | 521 | Notation "\sum_ ( m <= i < n ) F" := 522 | (foldr (fun i a => F + a) 0 (iota m (n-m))). 523 | 524 | Check \sum_(1 <= x < 5) (x * 2 - 1). 525 | Eval lazy in \sum_(1 <= x < 5) (x * 2 - 1). 526 | (** 527 | #
# 528 | 529 | #


# 530 | 531 | #

(notes)
# 532 | 533 | This slide corresponds to 534 | section 1.6 of 535 | #the Mathematical Components book# 536 | #
# 537 | 538 | #
# 539 | 540 | ---------------------------------------------------------- 541 | #
# 542 | ** Lesson 1: sum up 543 | 544 | - [fun .. => ..] 545 | - [Check] 546 | - [Definition] 547 | - [Print] 548 | - [Eval lazy] 549 | - [Inductive] declarations [bool], [nat], [seq]. 550 | - [match .. with .. end] and [if .. is .. then .. else ..] 551 | - [Fixpoint] 552 | - [andb] [orb] [eqn] [leq] [addn] [subn] [size] [foldr] 553 | 554 | #
# 555 | 556 | 557 | *) 558 | --------------------------------------------------------------------------------