completed pass over Quantifiers
This commit is contained in:
parent
d878b55f2e
commit
8b7519220f
6 changed files with 186 additions and 105 deletions
|
@ -694,7 +694,7 @@ is the same as the assertion that if `A` holds then `B` holds and if
|
|||
`A` holds then `C` holds. The proof of left inverse requires both extensionality
|
||||
and the rule `η-×` for products.
|
||||
\begin{code}
|
||||
→-distrib-× : ∀ {A B C : Set} → (A → B × C) ≃ ((A → B) × (A → C))
|
||||
→-distrib-× : ∀ {A B C : Set} → (A → B × C) ≃ (A → B) × (A → C)
|
||||
→-distrib-× =
|
||||
record
|
||||
{ to = λ{ f → ( (λ{ x → proj₁ (f x) }) , (λ{ y → proj₂ (f y)}) ) }
|
||||
|
@ -710,39 +710,39 @@ and the rule `η-×` for products.
|
|||
Products distributes over sum, up to isomorphism. The code to validate
|
||||
this fact is similar in structure to our previous results.
|
||||
\begin{code}
|
||||
×-distrib-⊎ : ∀ {A B C : Set} → ((A ⊎ B) × C) ≃ ((A × C) ⊎ (B × C))
|
||||
×-distrib-⊎ : ∀ {A B C : Set} → (A ⊎ B) × C ≃ (A × C) ⊎ (B × C)
|
||||
×-distrib-⊎ =
|
||||
record
|
||||
{ to = λ { ((inj₁ x) , z) → (inj₁ (x , z))
|
||||
; ((inj₂ y) , z) → (inj₂ (y , z))
|
||||
}
|
||||
; from = λ { (inj₁ (x , z)) → ((inj₁ x) , z)
|
||||
; (inj₂ (y , z)) → ((inj₂ y) , z)
|
||||
}
|
||||
; from∘to = λ { ((inj₁ x) , z) → refl
|
||||
; ((inj₂ y) , z) → refl
|
||||
}
|
||||
; to∘from = λ { (inj₁ (x , z)) → refl
|
||||
; (inj₂ (y , z)) → refl
|
||||
}
|
||||
{ to = λ{ ((inj₁ x) , z) → (inj₁ (x , z))
|
||||
; ((inj₂ y) , z) → (inj₂ (y , z))
|
||||
}
|
||||
; from = λ{ (inj₁ (x , z)) → ((inj₁ x) , z)
|
||||
; (inj₂ (y , z)) → ((inj₂ y) , z)
|
||||
}
|
||||
; from∘to = λ{ ((inj₁ x) , z) → refl
|
||||
; ((inj₂ y) , z) → refl
|
||||
}
|
||||
; to∘from = λ{ (inj₁ (x , z)) → refl
|
||||
; (inj₂ (y , z)) → refl
|
||||
}
|
||||
}
|
||||
\end{code}
|
||||
|
||||
Sums do not distribute over products up to isomorphism, but it is an embedding.
|
||||
\begin{code}
|
||||
⊎-distrib-× : ∀ {A B C : Set} → ((A × B) ⊎ C) ≲ ((A ⊎ C) × (B ⊎ C))
|
||||
⊎-distrib-× : ∀ {A B C : Set} → (A × B) ⊎ C ≲ (A ⊎ C) × (B ⊎ C)
|
||||
⊎-distrib-× =
|
||||
record
|
||||
{ to = λ { (inj₁ (x , y)) → (inj₁ x , inj₁ y)
|
||||
; (inj₂ z) → (inj₂ z , inj₂ z)
|
||||
}
|
||||
; from = λ { (inj₁ x , inj₁ y) → (inj₁ (x , y))
|
||||
; (inj₁ x , inj₂ z) → (inj₂ z)
|
||||
; (inj₂ z , _ ) → (inj₂ z)
|
||||
}
|
||||
; from∘to = λ { (inj₁ (x , y)) → refl
|
||||
; (inj₂ z) → refl
|
||||
}
|
||||
{ to = λ{ (inj₁ (x , y)) → (inj₁ x , inj₁ y)
|
||||
; (inj₂ z) → (inj₂ z , inj₂ z)
|
||||
}
|
||||
; from = λ{ (inj₁ x , inj₁ y) → (inj₁ (x , y))
|
||||
; (inj₁ x , inj₂ z) → (inj₂ z)
|
||||
; (inj₂ z , _ ) → (inj₂ z)
|
||||
}
|
||||
; from∘to = λ{ (inj₁ (x , y)) → refl
|
||||
; (inj₂ z) → refl
|
||||
}
|
||||
}
|
||||
\end{code}
|
||||
Note that there is a choice in how we write the `from` function.
|
||||
|
@ -763,6 +763,16 @@ gives rise to an isomorphism, while the second only gives rise to an
|
|||
embedding, revealing a sense in which one of these laws is "more
|
||||
true" than the other.
|
||||
|
||||
|
||||
### Exercise (`×⊎-implies-⊎×`)
|
||||
|
||||
Show that a conjunct of disjuncts implies a disjunct of conjuncts.
|
||||
\begin{code}
|
||||
×⊎-Implies-⊎× = ∀ {A B C D : Set} → (A ⊎ B) × (C ⊎ D) → (A × C) ⊎ (B × D)
|
||||
\end{code}
|
||||
Does the converse hold? If so, prove; if not, explain why.
|
||||
|
||||
|
||||
### Exercise (`⇔-refl`, `⇔-sym`, `⇔-trans`)
|
||||
|
||||
Define equivalence of propositions (also known as "if and only if") as follows.
|
||||
|
|
|
@ -35,6 +35,7 @@ then applies `g`.
|
|||
In set theory, two sets are isomorphic if they are in one-to-one correspondence.
|
||||
Here is a formal definition of isomorphism.
|
||||
\begin{code}
|
||||
infix 0 _≃_
|
||||
record _≃_ (A B : Set) : Set where
|
||||
field
|
||||
to : A → B
|
||||
|
@ -216,6 +217,7 @@ correspondence between the second type and the first.
|
|||
|
||||
Here is the formal definition of embedding.
|
||||
\begin{code}
|
||||
infix 0 _≲_
|
||||
record _≲_ (A B : Set) : Set where
|
||||
field
|
||||
to : A → B
|
||||
|
|
|
@ -93,7 +93,7 @@ Show that a disjunction of universals implies a universal of disjunctions.
|
|||
⊎∀-Implies-∀⊎ = ∀ {A : Set} { B C : A → Set } →
|
||||
(∀ (x : A) → B x) ⊎ (∀ (x : A) → C x) → ∀ (x : A) → B x ⊎ C x
|
||||
\end{code}
|
||||
Does the converse also hold? If so, prove; if not, explain why.
|
||||
Does the converse hold? If so, prove; if not, explain why.
|
||||
|
||||
|
||||
## Existentials
|
||||
|
@ -133,24 +133,29 @@ are provided by the evidence for `∃ (λ (x : A) → B x)`.
|
|||
|
||||
Agda makes it possible to define our own syntactic abbreviations.
|
||||
\begin{code}
|
||||
syntax ∃ {A} (λ x → B) = ∃[ x ∈ A ] B
|
||||
syntax ∃ (λ x → B) = ∃[ x ] B
|
||||
\end{code}
|
||||
This allows us to write `∃[ x ∈ A ] (B x)` in place of
|
||||
`∃ (λ (x : A) → B x)`.
|
||||
This allows us to write `∃[ x ] (B x)` in place of `∃ (λ x → B x)`.
|
||||
|
||||
As an example, recall the definitions of `even` from
|
||||
As an example, recall the definitions of `even` and `odd` from
|
||||
Chapter [Relations](Relations).
|
||||
\begin{code}
|
||||
data even : ℕ → Set where
|
||||
ev0 : even zero
|
||||
ev+2 : ∀ {n : ℕ} → even n → even (suc (suc n))
|
||||
data even : ℕ → Set
|
||||
data odd : ℕ → Set
|
||||
|
||||
data even where
|
||||
even-zero : even zero
|
||||
even-suc : ∀ {n : ℕ} → odd n → even (suc n)
|
||||
|
||||
data odd where
|
||||
odd-suc : ∀ {n : ℕ} → even n → odd (suc n)
|
||||
\end{code}
|
||||
A number is even if it is zero, or if it is two
|
||||
greater than an even number.
|
||||
A number is even if it is zero or the successor of an odd number, and
|
||||
odd if it the successor of an even number.
|
||||
|
||||
We will show that a number is even if and only if it is twice some
|
||||
other number. That is, number `n` is even if and only if there exists
|
||||
a number `m` such that twice `m` is `n`.
|
||||
other number, and odd if and only if it is one more than twice
|
||||
some other number.
|
||||
|
||||
First, we need a lemma, which allows us to
|
||||
simplify twice the successor of `m` to two
|
||||
|
@ -175,48 +180,71 @@ The lemma is straightforward, and uses the lemma
|
|||
allows us to simplify `m + suc n` to `suc (m + n)`.
|
||||
|
||||
Here is the proof in the forward direction.
|
||||
\begin{code}
|
||||
ev-ex : ∀ {n : ℕ} → even n → ∃[ m ∈ ℕ ] (2 * m ≡ n)
|
||||
ev-ex ev0 = (zero , refl)
|
||||
ev-ex (ev+2 ev) with ev-ex ev
|
||||
... | (m , refl) = (suc m , lemma m)
|
||||
\begin{code}
|
||||
even-∃ : ∀ {n : ℕ} → even n → ∃[ m ] ( 2 * m ≡ n)
|
||||
odd-∃ : ∀ {n : ℕ} → odd n → ∃[ m ] (1 + 2 * m ≡ n)
|
||||
|
||||
even-∃ even-zero = zero , refl
|
||||
even-∃ (even-suc o) with odd-∃ o
|
||||
... | m , refl = suc m , lemma m
|
||||
|
||||
odd-∃ (odd-suc e) with even-∃ e
|
||||
... | m , refl = m , refl
|
||||
\end{code}
|
||||
Given an even number, we must show it is twice some
|
||||
other number. The proof is a function, which
|
||||
given a proof that `n` is even
|
||||
returns a pair consisting of `m` and a proof
|
||||
that twice `m` is `n`. The proof is by induction over
|
||||
the evidence that `n` is even.
|
||||
We define two mutually recursive functions. Given
|
||||
evidence that `n` is even or odd, we return a
|
||||
number `m` and evidence that `2 * m ≡ n` or `1 + 2 * m ≡ n`.
|
||||
(By convention, one tends to put a constant at the
|
||||
end of a term, so why have we chosen here
|
||||
to write `1 + 2 * m` rather than `2 * m + 1`?)
|
||||
We induct over the evidence that `n` is even or odd.
|
||||
|
||||
- If the number is even because it is zero, then we return a pair
|
||||
consisting of zero and the (trivial) proof that twice zero is zero.
|
||||
|
||||
- If the number is even because it is two more than another even
|
||||
number `n`, then we apply the induction hypothesis, giving us a number
|
||||
`m` and a proof that `2 * m ≡ n`, which we match against `refl`. We
|
||||
return a pair consisting of `suc m` and a proof that `2 * suc m ≡ suc
|
||||
(suc n)`, which follows from `2 * m ≡ n` and the lemma.
|
||||
- If the number is even because it is one more than an odd number,
|
||||
then we apply the induction hypothesis to give a number `m` and
|
||||
evidence that `1 + 2 * m ≡ n`. We return a pair consisting of `suc m`
|
||||
and evidence that `2 * (suc m)` ≡ suc n`, which after substituting for
|
||||
`n` means we need to show `2 * (suc m) ≡ 2 + 2 * m`, which follows
|
||||
from our lemma.
|
||||
|
||||
- If the number is odd because it is the successor of an even
|
||||
number, then we apply the induction hypothesis to give a number `m`
|
||||
and evidence that `2 * m ≡ n`. We return a pair conisting of `suc m`
|
||||
and evidence that `1 + 2 * m = suc n`, which is immediate
|
||||
after substituting for `n`.
|
||||
|
||||
This completes the proof in the forward direction.
|
||||
|
||||
Here is the proof in the reverse direction.
|
||||
\begin{code}
|
||||
ex-ev : ∀ {n : ℕ} → ∃[ m ∈ ℕ ] (2 * m ≡ n) → even n
|
||||
ex-ev (zero , refl) = ev0
|
||||
ex-ev (suc m , refl) rewrite lemma m = ev+2 (ex-ev (m , refl))
|
||||
∃-even : ∀ {n : ℕ} → ∃[ m ] ( 2 * m ≡ n) → even n
|
||||
∃-odd : ∀ {n : ℕ} → ∃[ m ] (1 + 2 * m ≡ n) → odd n
|
||||
|
||||
∃-even ( zero , refl) = even-zero
|
||||
∃-even (suc m , refl) rewrite lemma m = even-suc (∃-odd (m , refl))
|
||||
|
||||
∃-odd ( m , refl) = odd-suc (∃-even (m , refl))
|
||||
\end{code}
|
||||
Given a number that is twice some other number,
|
||||
we must show that it is even. The proof is a function,
|
||||
which given a number `m` and a proof that `n` is twice `m`,
|
||||
returns a proof that `n` is even. The proof is by induction
|
||||
over the number `m`.
|
||||
Given a number that is twice some other number we must show
|
||||
it is even, and a number that is one more than twice some other
|
||||
number we must show it is odd. We induct over the evidence
|
||||
of the existential, and in particular the number that is doubled.
|
||||
|
||||
- If it is zero, then we must show that twice zero is even,
|
||||
which follows by rule `ev0`.
|
||||
- In the even case, if it is `zero`, then we must show `2 * zero` is
|
||||
even, which follows by `even-zero`.
|
||||
|
||||
- If it is `suc m`, then we must show that `2 * suc m` is
|
||||
even. After rewriting with our lemma, we must show that
|
||||
`suc (suc (2 * m))` is even. The inductive hypothesis tells
|
||||
us `2 * m` is even, from which the desired result follows
|
||||
by rule `ev+2`.
|
||||
- In the even case, if it is `suc n`, then we must show `2 * suc m` is
|
||||
even. After rewriting with our lemma, we must show that `2 + 2 * m`
|
||||
is even. The inductive hypothesis tells us that `1 + 2 * m` is odd,
|
||||
from which the desired result follows by `even-suc`.
|
||||
|
||||
- In the odd case, then we must show `1 + 2 * suc m` is odd. The
|
||||
inductive hypothesis tell us that `2 * m` is even, from which the
|
||||
desired result follows by `odd-suc`.
|
||||
|
||||
This completes the proof in the backward direction.
|
||||
|
||||
Negation of an existential is isomorphic to universal
|
||||
of a negation. Considering that existentials are generalised
|
||||
|
@ -224,44 +252,67 @@ disjuntion and universals are generalised conjunction, this
|
|||
result is analogous to the one which tells us that negation
|
||||
of a disjuntion is isomorphic to a conjunction of negations.
|
||||
\begin{code}
|
||||
¬∃∀ : ∀ {A : Set} {B : A → Set} → (¬ ∃[ x ∈ A ] B x) ≃ ∀ (x : A) → ¬ B x
|
||||
¬∃∀ : ∀ {A : Set} {B : A → Set} → (¬ ∃[ x ] B x) ≃ ∀ x → ¬ B x
|
||||
¬∃∀ =
|
||||
record
|
||||
{ to = λ { ¬∃bx x bx → ¬∃bx (x , bx) }
|
||||
; from = λ { ∀¬bx (x , bx) → ∀¬bx x bx }
|
||||
; from∘to = λ { ¬∃bx → extensionality (λ { (x , bx) → refl }) }
|
||||
; to∘from = λ { ∀¬bx → refl }
|
||||
{ to = λ{ ¬∃xy x y → ¬∃xy (x , y) }
|
||||
; from = λ{ ∀¬xy (x , y) → ∀¬xy x y }
|
||||
; from∘to = λ{ ¬∃xy → extensionality λ{ (x , y) → refl } }
|
||||
; to∘from = λ{ ∀¬xy → refl }
|
||||
}
|
||||
\end{code}
|
||||
In the `to` direction, we are given a value `¬∃bx` of type
|
||||
`¬ ∃ (λ (x : A) → B x)`, and need to show that given a value
|
||||
`x` of type `A` that `¬ B x` follows, in other words, from
|
||||
a value `bx` of type `B x` we can derive false. Combining
|
||||
`x` and `bx` gives us a value `(x , bx)` of type `∃ (λ (x : A) → B x)`,
|
||||
and applying `¬∃bx` to that yields a contradiction.
|
||||
In the `to` direction, we are given a value `¬∃xy` of type
|
||||
`¬ ∃[ x ] B x`, and need to show that given a value
|
||||
`x` that `¬ B x` follows, in other words, from
|
||||
a value `y` of type `B x` we can derive false. Combining
|
||||
`x` and `y` gives us a value `(x , y)` of type `∃[ x ] B x`,
|
||||
and applying `¬∃xy` to that yields a contradiction.
|
||||
|
||||
In the `from` direction, we are given a value `∀¬bx` of type
|
||||
`∀ (x : A) → ¬ B x`, and need to show that from a value `(x , bx)`
|
||||
of type `∃ (λ (x : A) → B x)` we can derive false. Applying `∀¬bx`
|
||||
to `x` gives a value of type `¬ B x`, and applying that to `bx` yields
|
||||
In the `from` direction, we are given a value `∀¬xy` of type
|
||||
`∀ x → ¬ B x`, and need to show that from a value `(x , y)`
|
||||
of type `∃[ x ] B x` we can derive false. Applying `∀¬xy`
|
||||
to `x` gives a value of type `¬ B x`, and applying that to `y` yields
|
||||
a contradiction.
|
||||
|
||||
The two inverse proofs are straightforward, where one direction
|
||||
requires extensionality.
|
||||
|
||||
*Exercise* Show `∃ (λ (x : A) → ¬ B x) → ¬ (∀ (x : A) → B x)`.
|
||||
### Exercise (`∃-distrib-⊎`)
|
||||
|
||||
Show that universals distribute over conjunction.
|
||||
\begin{code}
|
||||
∃-Distrib-⊎ = ∀ {A : Set} {B C : A → Set} →
|
||||
∃[ x ] (B x ⊎ C x) ≃ (∃[ x ] B x) ⊎ (∃[ x ] C x)
|
||||
\end{code}
|
||||
|
||||
### Exercise (`∃×-implies-×∃`)
|
||||
|
||||
Show that an existential of conjunctions implies a conjunction of existentials.
|
||||
\begin{code}
|
||||
∃×-Implies-×∃ = ∀ {A : Set} { B C : A → Set } →
|
||||
∃[ x ] (B x × C x) → (∃[ x ] B x) × (∃[ x ] C x)
|
||||
\end{code}
|
||||
Does the converse hold? If so, prove; if not, explain why.
|
||||
|
||||
### Exercise (`∃¬-Implies-¬∀)
|
||||
|
||||
Show `∃[ x ] ¬ B x → ¬ (∀ x → B x)`.
|
||||
|
||||
Does the converse hold? If so, prove; if not, explain why.
|
||||
|
||||
|
||||
## Standard Prelude
|
||||
|
||||
Definitions similar to those in this chapter can be found in the standard library.
|
||||
\begin{code}
|
||||
import Data.Product using (∃)
|
||||
import Data.Product using (∃;_,_)
|
||||
\end{code}
|
||||
|
||||
|
||||
## Unicode
|
||||
|
||||
This chapter introduces the following unicode.
|
||||
This chapter uses the following unicode.
|
||||
|
||||
≤ U+2264 LESS-THAN OR EQUAL TO (\<=, \le)
|
||||
∃ U+2203 THERE EXISTS (\ex, \exists)
|
||||
|
||||
|
||||
|
|
|
@ -540,7 +540,7 @@ data odd where
|
|||
odd-suc : ∀ {n : ℕ} → even n → odd (suc n)
|
||||
\end{code}
|
||||
A number is even if it is zero or the successor of an odd number,
|
||||
and a number is odd if it is the successor of an even number.
|
||||
and odd if it is the successor of an even number.
|
||||
|
||||
This is our first use of a mutually recursive datatype declaration.
|
||||
Since each identifier must be defined before it is used, we first
|
||||
|
|
|
@ -14,22 +14,16 @@ suc m ≤? suc n with m ≤? n
|
|||
... | yes m≤n = yes (s≤s m≤n)
|
||||
... | no ¬m≤n = no λ{ (s≤s m≤n) → ¬m≤n m≤n }
|
||||
|
||||
_ : Dec (2 ≤ 4)
|
||||
_ = 2 ≤? 4 -- yes (s≤s (s≤s z≤n))
|
||||
|
||||
_ : 2 ≤? 4 ≡ yes (s≤s (s≤s z≤n))
|
||||
_ = refl
|
||||
|
||||
_ : Dec (4 ≤ 2)
|
||||
_ = 4 ≤? 2 -- no λ{(s≤s (s≤s ()))}
|
||||
|
||||
_ : 4 ≤? 2 ≡ no (λ { (s≤s m≤n) → (λ { (s≤s m≤n) → (λ ()) m≤n }) m≤n })
|
||||
_ : 4 ≤? 2 ≡ no {!!}
|
||||
_ = refl
|
||||
|
||||
{-
|
||||
/Users/wadler/sf/src/extra/DecidableBroken.agda:27,5-9
|
||||
(λ { (s≤s m≤n) → (λ { (s≤s m≤n) → (λ ()) 1 m≤n }) m≤n }) x !=
|
||||
(λ { (s≤s m≤n) → _34 m≤n }) x of type .Data.Empty.⊥
|
||||
when checking that the expression refl has type
|
||||
(4 ≤? 2) ≡ no (λ { (s≤s m≤n) → _34 m≤n })
|
||||
Using ^C ^N, the term
|
||||
4 ≤? 2
|
||||
evaluates to
|
||||
no (λ { (s≤s m≤n) → (λ { (s≤s m≤n) → (λ ()) 1 m≤n }) m≤n })
|
||||
The 1 is spurious.
|
||||
-}
|
||||
|
|
|
@ -1,5 +1,6 @@
|
|||
import Relation.Binary.PropositionalEquality as Eq
|
||||
open Eq using (_≡_; refl; sym; trans; cong)
|
||||
open Eq.≡-Reasoning
|
||||
open import Data.Nat using (ℕ; zero; suc; _+_; _*_)
|
||||
open import Data.Nat.Properties.Simple using (+-suc)
|
||||
open import Data.Product using (∃; ∃-syntax; _,_)
|
||||
|
@ -14,23 +15,46 @@ data even where
|
|||
data odd where
|
||||
odd-suc : ∀ {n : ℕ} → even n → odd (suc n)
|
||||
|
||||
∃-even : ∀ {n : ℕ} → even n → ∃[ m ] (n ≡ 2 * m)
|
||||
∃-odd : ∀ {n : ℕ} → odd n → ∃[ m ] (n ≡ 1 + 2 * m)
|
||||
lemma : ∀ (m : ℕ) → 2 * suc m ≡ suc (suc (2 * m))
|
||||
lemma m =
|
||||
begin
|
||||
2 * suc m
|
||||
≡⟨⟩
|
||||
suc m + (suc m + zero)
|
||||
≡⟨⟩
|
||||
suc (m + (suc (m + zero)))
|
||||
≡⟨ cong suc (+-suc m (m + zero)) ⟩
|
||||
suc (suc (m + (m + zero)))
|
||||
≡⟨⟩
|
||||
suc (suc (2 * m))
|
||||
∎
|
||||
|
||||
∃-even : ∀ {n : ℕ} → even n → ∃[ m ] (2 * m ≡ n)
|
||||
∃-odd : ∀ {n : ℕ} → odd n → ∃[ m ] (1 + 2 * m ≡ n)
|
||||
|
||||
∃-even even-zero = zero , refl
|
||||
∃-even (even-suc o) with ∃-odd o
|
||||
... | m , eqn rewrite eqn | sym (+-suc m (m + 0)) = suc m , refl
|
||||
... | m , refl = suc m , lemma m
|
||||
|
||||
∃-odd (odd-suc e) with ∃-even e
|
||||
... | m , eqn rewrite eqn = m , refl
|
||||
... | m , refl = m , refl
|
||||
|
||||
∃-even′ : ∀ {n : ℕ} → even n → ∃[ m ] (n ≡ 2 * m)
|
||||
∃-odd′ : ∀ {n : ℕ} → odd n → ∃[ m ] (n ≡ 1 + 2 * m)
|
||||
|
||||
∃-even′ even-zero = zero , refl
|
||||
∃-even′ (even-suc o) with ∃-odd′ o
|
||||
... | m , eqn rewrite eqn | +-suc m (m + 0) = suc m , ?
|
||||
... | m , eqn rewrite eqn | +-suc m (m + 0) = suc m , {!!}
|
||||
|
||||
∃-odd′ (odd-suc e) with ∃-even′ e
|
||||
... | m , eqn rewrite eqn = m , refl
|
||||
|
||||
data Even : ℕ → Set where
|
||||
ev0 : Even zero
|
||||
ev2 : ∀ {n} → Even n → Even (suc (suc n))
|
||||
|
||||
ev-ex : ∀ {n : ℕ} → Even n → ∃[ m ] (2 * m ≡ n)
|
||||
ev-ex ev0 = (zero , refl)
|
||||
ev-ex (ev2 ev) with ev-ex ev
|
||||
... | (m , refl) = (suc m , lemma m)
|
||||
|
||||
|
|
Loading…
Reference in a new issue