<div dir="ltr">Thanks to Amr and Miëtek for their responses. Here is a follow up question.<div><br></div><div>First, is there a name in the standard library for the operation called <span style="color:rgb(34,34,34);font-family:arial,sans-serif;font-size:small;font-style:normal;font-variant-ligatures:normal;font-variant-caps:normal;font-weight:400;letter-spacing:normal;text-align:start;text-indent:0px;text-transform:none;white-space:normal;word-spacing:0px;background-color:rgb(255,255,255);text-decoration-style:initial;text-decoration-color:initial;float:none;display:inline">_≤ᵇ_ below?</span></div><div><br></div><div>Second, and less trivially, is there a way to fill in the holes in the proof of <span style="color:rgb(34,34,34);font-family:arial,sans-serif;font-size:small;font-style:normal;font-variant-ligatures:normal;font-variant-caps:normal;font-weight:400;letter-spacing:normal;text-align:start;text-indent:0px;text-transform:none;white-space:normal;word-spacing:0px;background-color:rgb(255,255,255);text-decoration-style:initial;text-decoration-color:initial;float:none;display:inline">_≤?′_ below, or to complete a proof along similar lines?</span></div><div><br></div><div>Many thanks, -- P<br><div><br></div></div><div><br></div><div><div>open import Data.Nat using (ℕ; zero; suc; _≤_; z≤n; s≤s)</div><div>open import Relation.Nullary using (¬_)</div><div>open import Relation.Nullary.Negation using (contraposition)</div><div>open import Data.Unit using (⊤; tt)</div><div>open import Data.Empty using (⊥)</div><div><br></div><div>data Bool : Set where</div><div>  true : Bool</div><div>  false : Bool</div><div><br></div><div>T : Bool → Set</div><div>T true = ⊤</div><div>T false = ⊥</div><div><br></div><div>_≤ᵇ_ : ℕ → ℕ → Bool</div><div>zero ≤ᵇ n = true</div><div>suc m ≤ᵇ zero = false</div><div>suc m ≤ᵇ suc n = m ≤ᵇ n</div><div><br></div><div>≤→≤ᵇ : ∀ {m n : ℕ} → m ≤ n → T (m ≤ᵇ n)</div><div>≤→≤ᵇ z≤n = tt</div><div>≤→≤ᵇ (s≤s m≤n) = ≤→≤ᵇ m≤n</div><div><br></div><div>≤ᵇ→≤ : ∀ (m n : ℕ) → T (m ≤ᵇ n) → m ≤ n</div><div>≤ᵇ→≤ zero n tt = z≤n</div><div>≤ᵇ→≤ (suc m) zero ()</div><div>≤ᵇ→≤ (suc m) (suc n) m≤ᵇn =  s≤s (≤ᵇ→≤ m n m≤ᵇn)</div><div><br></div><div>data Dec (A : Set) : Set where</div><div>  yes : A → Dec A</div><div>  no : ¬ A → Dec A</div><div><br></div><div>_≤?_ : ∀ (m n : ℕ) → Dec (m ≤ n)</div><div>zero ≤? n = yes z≤n</div><div>suc m ≤? zero = no λ()</div><div>suc m ≤? suc n with m ≤? n</div><div>... | yes m≤n = yes (s≤s m≤n)</div><div>... | no ¬m≤n = no λ{ (s≤s m≤n) → ¬m≤n m≤n }</div><div><br></div><div>_≤?′_ : ∀ (m n : ℕ) → Dec (m ≤ n)</div><div>m ≤?′ n with m ≤ᵇ n</div><div>... | true = yes (≤ᵇ→≤ m n {!!})</div><div>... | false = no (contraposition ≤→≤ᵇ {!!})</div></div><div><br></div></div><div class="gmail_extra"><br clear="all"><div><div class="gmail_signature" data-smartmail="gmail_signature"><div dir="ltr"><div><div dir="ltr">.   \ Philip Wadler, Professor of Theoretical Computer Science,<br>.   /\ School of Informatics, University of Edinburgh<br></div><div>.  /  \ and Senior Research Fellow, IOHK<br></div><div dir="ltr">. <span><a href="http://homepages.inf.ed.ac.uk/wadler/" target="_blank">http://homepages.inf.ed.ac.uk/wadler/</a></span></div></div></div></div></div>
<br><div class="gmail_quote">On 15 February 2018 at 18:35, Miëtek Bak <span dir="ltr"><<a href="mailto:mietek@bak.io" target="_blank">mietek@bak.io</a>></span> wrote:<br><blockquote class="gmail_quote" style="margin:0 0 0 .8ex;border-left:1px #ccc solid;padding-left:1ex"><div style="word-wrap:break-word"><div>The standard library defines synonyms that obscure the underlying types.  I haven’t found C-c C-z to be of much use; I grep the source trees instead.</div><div><br></div><div>$ git grep Dec | grep Nat</div><div>src/Data/Nat/Base.agda:156:_≤?<wbr>_ : Decidable _≤_</div><div><br></div><div>Here’s the second operation:</div><div><a href="https://agda.github.io/agda-stdlib/Data.Nat.Base.html#3179" target="_blank">https://agda.github.io/agda-<wbr>stdlib/Data.Nat.Base.html#3179</a></div><div><br></div><div>We can obtain the first operation via projection:</div><div><a href="https://agda.github.io/agda-stdlib/Relation.Nullary.Decidable.html#822" target="_blank">https://agda.github.io/agda-<wbr>stdlib/Relation.Nullary.<wbr>Decidable.html#822</a></div><span class="HOEnZb"><font color="#888888"><div><br></div><div><br></div><div>-- </div><div>M.</div><div><br></div><div><br></div></font></span><div><br><div><blockquote type="cite"><div><div class="h5"><div>On 15 Feb 2018, at 20:26, Philip Wadler <<a href="mailto:wadler@inf.ed.ac.uk" target="_blank">wadler@inf.ed.ac.uk</a>> wrote:</div><br class="m_2941185566432368495Apple-interchange-newline"></div></div><div><div><div class="h5"><div dir="ltr"><div>I presume there are operations in the standard prelude to compute whether one natural is less than or equal to another with types</div><div><br></div><div>  Nat -> Nat -> Bool</div><div>  (n : Nat) -> (m : Nat) -> Dec (m \leq n)</div><div><br></div><div>But I can't find them. Where are they and what are they called? Cheers, -- P</div><div><br></div><div>PS. Using ^C ^Z to search Everything.agda (see previous thread) for Nat Bool or Nat Dec yields nothing. Indeed, using it to search for Nat gives:</div><div><br></div><div>  Definitions about Nat</div><div><br></div><div>followed by nothing! I'm not sure why.</div><div><br></div><div> </div><div><br></div><br clear="all"><div><div class="m_2941185566432368495gmail_signature"><div dir="ltr"><div><div dir="ltr">.   \ Philip Wadler, Professor of Theoretical Computer Science,<br>.   /\ School of Informatics, University of Edinburgh<br></div><div>.  /  \ and Senior Research Fellow, IOHK<br></div><div dir="ltr">. <span><a href="http://homepages.inf.ed.ac.uk/wadler/" target="_blank">http://homepages.inf.ed.ac.uk/<wbr>wadler/</a></span></div></div></div></div></div>
</div></div></div><span class="">
The University of Edinburgh is a charitable body, registered in<br>Scotland, with registration number SC005336.<br>______________________________<wbr>_________________<br>Agda mailing list<br><a href="mailto:Agda@lists.chalmers.se" target="_blank">Agda@lists.chalmers.se</a><br><a href="https://lists.chalmers.se/mailman/listinfo/agda" target="_blank">https://lists.chalmers.se/<wbr>mailman/listinfo/agda</a><br></span></div></blockquote></div><br></div></div></blockquote></div><br></div>