Skip to content

Commit a0f3343

Browse files
authored
Add cong! tactic (#1658)
1 parent 416904e commit a0f3343

File tree

7 files changed

+378
-8
lines changed

7 files changed

+378
-8
lines changed

CHANGELOG.md

+5
Original file line numberDiff line numberDiff line change
@@ -564,6 +564,11 @@ New modules
564564
Reflection.AlphaEquality
565565
```
566566

567+
* `cong!` tactic for deriving arguments to `cong`
568+
```
569+
Tactic.Rewrite
570+
```
571+
567572
* Various system types and primitives:
568573
```
569574
System.Clock.Primitive

README/Tactic/Rewrite.agda

+139
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,139 @@
1+
{-# OPTIONS --without-K --safe #-}
2+
module README.Tactic.Rewrite where
3+
4+
open import Data.Nat
5+
open import Data.Nat.Properties
6+
7+
open import Relation.Binary.PropositionalEquality as Eq
8+
import Relation.Binary.Reasoning.Preorder as PR
9+
10+
open import Tactic.Rewrite using (cong!)
11+
12+
----------------------------------------------------------------------
13+
-- Usage
14+
----------------------------------------------------------------------
15+
16+
-- When performing large equational reasoning proofs, it's quite
17+
-- common to have to construct sophisticated lambdas to pass
18+
-- into 'cong'. This can be extremely tedious, and can bog down
19+
-- large proofs in piles of boilerplate. The 'cong!' tactic
20+
-- simplifies this process by synthesizing the appropriate call
21+
-- to 'cong' by inspecting both sides of the goal.
22+
--
23+
-- This is best demonstrated with a small example. Consider
24+
-- the following proof:
25+
26+
verbose-example : m n m ≡ n suc (suc (m + 0)) + m ≡ suc (suc n) + (n + 0)
27+
verbose-example m n eq =
28+
let open Eq.≡-Reasoning in
29+
begin
30+
suc (suc (m + 0)) + m
31+
≡⟨ cong (λ ϕ suc (suc (ϕ + m))) (+-identityʳ m) ⟩
32+
suc (suc m) + m
33+
≡⟨ cong (λ ϕ suc (suc (ϕ + ϕ))) eq ⟩
34+
suc (suc n) + n
35+
≡˘⟨ cong (λ ϕ suc (suc (n + ϕ))) (+-identityʳ n) ⟩
36+
suc (suc n) + (n + 0)
37+
38+
39+
-- The calls to 'cong' add a lot of boilerplate, and also
40+
-- clutter up the proof, making it more difficult to read.
41+
-- We can simplify this by using 'cong!' to deduce those
42+
-- lambdas for us.
43+
44+
succinct-example : m n m ≡ n suc (suc (m + 0)) + m ≡ suc (suc n) + (n + 0)
45+
succinct-example m n eq =
46+
let open Eq.≡-Reasoning in
47+
begin
48+
suc (suc (m + 0)) + m
49+
≡⟨ cong! (+-identityʳ m) ⟩
50+
suc (suc m) + m
51+
≡⟨ cong! eq ⟩
52+
suc (suc n) + n
53+
≡˘⟨ cong! (+-identityʳ n) ⟩
54+
suc (suc n) + (n + 0)
55+
56+
57+
----------------------------------------------------------------------
58+
-- Limitations
59+
----------------------------------------------------------------------
60+
61+
-- The 'cong!' tactic can handle simple cases, but will
62+
-- struggle when presented with equality proofs like
63+
-- 'm + n ≡ n + m' or 'm + (n + o) ≡ (m + n) + o'.
64+
--
65+
-- The reason behind this is that this tactic operates by simple
66+
-- anti-unification; it examines both sides of the equality goal
67+
-- to deduce where to generalize. When presented with two sides
68+
-- of an equality like 'm + n ≡ n + m', it will anti-unify to
69+
-- 'ϕ + ϕ', which is too specific.
70+
71+
----------------------------------------------------------------------
72+
-- Unit Tests
73+
----------------------------------------------------------------------
74+
75+
module LiteralTests
76+
(assumption : 4842)
77+
(f : ℕ)
78+
where
79+
80+
test₁ : 40 + 242
81+
test₁ = cong! refl
82+
83+
test₂ : 4842 4248
84+
test₂ eq = cong! (sym eq)
85+
86+
test₃ : (f : ℕ) f 48 ≡ f 42
87+
test₃ f = cong! assumption
88+
89+
test₄ : (f : ℕ) f 48 48 ≡ f 42 42
90+
test₄ f = cong! assumption
91+
92+
test₅ : f 48 45 48 ≡ f 42 45 42
93+
test₅ = cong! assumption
94+
95+
module LambdaTests
96+
(assumption : 4842)
97+
where
98+
99+
test₁ : (λ x x + 48) ≡ (λ x x + 42)
100+
test₁ = cong! assumption
101+
102+
test₂ : (λ x y z x + (y + 48 + z)) ≡ (λ x y z x + (y + 42 + z))
103+
test₂ = cong! assumption
104+
105+
module HigherOrderTests
106+
(f g : ℕ)
107+
where
108+
109+
test₁ : f ≡ g n f n ≡ g n
110+
test₁ eq n = cong! eq
111+
112+
test₂ : f ≡ g n f (f (f n)) ≡ g (g (g n))
113+
test₂ eq n = cong! eq
114+
115+
module EquationalReasoningTests where
116+
117+
test₁ : m n m ≡ n suc (suc (m + 0)) + m ≡ suc (suc n) + (n + 0)
118+
test₁ m n eq =
119+
let open Eq.≡-Reasoning in
120+
begin
121+
suc (suc (m + 0)) + m
122+
≡⟨ cong! (+-identityʳ m) ⟩
123+
suc (suc m) + m
124+
≡⟨ cong! eq ⟩
125+
suc (suc n) + n
126+
≡˘⟨ cong! (+-identityʳ n) ⟩
127+
suc (suc n) + (n + 0)
128+
129+
130+
test₂ : m n m ≡ n suc (m + m) ≤ suc (suc (n + n))
131+
test₂ m n eq =
132+
let open PR ≤-preorder in
133+
begin
134+
suc (m + m)
135+
≡⟨ cong! eq ⟩
136+
suc (n + n)
137+
∼⟨ ≤-step ≤-refl ⟩
138+
suc (suc (n + n))
139+

src/Reflection/AlphaEquality.agda

+6-6
Original file line numberDiff line numberDiff line change
@@ -106,7 +106,7 @@ mutual
106106
(arg i a) =α=-ArgPattern (arg i′ a′) = a =α=-Pattern a′
107107

108108
_=α=-Term_ : Term Term Bool
109-
(var x args) =α=-Term (var x′ args′) = (x ≡ᵇ x′) ∧ (args =α=-ArgsTerm args′)
109+
(var x args) =α=-Term (var x′ args′) = (x ℕ.≡ᵇ x′) ∧ (args =α=-ArgsTerm args′)
110110
(con c args) =α=-Term (con c′ args′) = (c =α= c′) ∧ (args =α=-ArgsTerm args′)
111111
(def f args) =α=-Term (def f′ args′) = (f =α= f′) ∧ (args =α=-ArgsTerm args′)
112112
(meta x args) =α=-Term (meta x′ args′) = (x =α= x′) ∧ (args =α=-ArgsTerm args′)
@@ -210,10 +210,10 @@ mutual
210210

211211
_=α=-Sort_ : Sort Sort Bool
212212
(set t ) =α=-Sort (set t′ ) = t =α=-Term t′
213-
(lit n ) =α=-Sort (lit n′ ) = n ≡ᵇ n′
213+
(lit n ) =α=-Sort (lit n′ ) = n ℕ.≡ᵇ n′
214214
(prop t ) =α=-Sort (prop t′ ) = t =α=-Term t′
215-
(propLit n) =α=-Sort (propLit n′) = n ≡ᵇ n′
216-
(inf n ) =α=-Sort (inf n′ ) = n ≡ᵇ n′
215+
(propLit n) =α=-Sort (propLit n′) = n ℕ.≡ᵇ n′
216+
(inf n ) =α=-Sort (inf n′ ) = n ℕ.≡ᵇ n′
217217
(unknown ) =α=-Sort (unknown ) = true
218218

219219
(set _ ) =α=-Sort (lit _ ) = false
@@ -267,11 +267,11 @@ mutual
267267

268268
_=α=-Pattern_ : Pattern Pattern Bool
269269
(con c ps) =α=-Pattern (con c′ ps′) = (c =α= c′) ∧ (ps =α=-ArgsPattern ps′)
270-
(var x ) =α=-Pattern (var x′ ) = x ≡ᵇ x′
270+
(var x ) =α=-Pattern (var x′ ) = x ℕ.≡ᵇ x′
271271
(lit l ) =α=-Pattern (lit l′ ) = l =α= l′
272272
(proj a ) =α=-Pattern (proj a′ ) = a =α= a′
273273
(dot t ) =α=-Pattern (dot t′ ) = t =α=-Term t′
274-
(absurd x) =α=-Pattern (absurd x′ ) = x ≡ᵇ x′
274+
(absurd x) =α=-Pattern (absurd x′ ) = x ℕ.≡ᵇ x′
275275

276276
(con x x₁) =α=-Pattern (dot x₂ ) = false
277277
(con x x₁) =α=-Pattern (var x₂ ) = false

src/Reflection/Literal.agda

+4
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@
88

99
module Reflection.Literal where
1010

11+
open import Data.Bool.Base as Bool using (Bool; true; false)
1112
import Data.Char as Char
1213
import Data.Float as Float
1314
import Data.Nat as ℕ
@@ -102,3 +103,6 @@ meta x ≟ char x₁ = no (λ ())
102103
meta x ≟ string x₁ = no (λ ())
103104
meta x ≟ name x₁ = no (λ ())
104105
meta x ≟ meta x₁ = Dec.map′ (cong meta) meta-injective (x Meta.≟ x₁)
106+
107+
_≡ᵇ_ : Literal Literal Bool
108+
l ≡ᵇ l′ = Dec.isYes (l ≟ l′)

src/Reflection/Meta.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -16,7 +16,7 @@ import Relation.Binary.Construct.On as On
1616
open import Relation.Binary.PropositionalEquality
1717

1818
open import Agda.Builtin.Reflection public
19-
using (Meta) renaming (primMetaToNat to toℕ)
19+
using (Meta) renaming (primMetaToNat to toℕ; primMetaEquality to _≡ᵇ_)
2020

2121
open import Agda.Builtin.Reflection.Properties public
2222
renaming (primMetaToNatInjective to toℕ-injective)

src/Reflection/Name.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -21,7 +21,7 @@ open import Relation.Binary.PropositionalEquality
2121
-- Re-export built-ins
2222

2323
open import Agda.Builtin.Reflection public
24-
using (Name) renaming (primQNameToWord64s to toWords)
24+
using (Name) renaming (primQNameToWord64s to toWords; primQNameEquality to _≡ᵇ_)
2525

2626
open import Agda.Builtin.Reflection.Properties public
2727
renaming (primQNameToWord64sInjective to toWords-injective)

0 commit comments

Comments
 (0)