<div dir="ltr">I have an Agda formalisation of reduction on raw terms (inspired by a formalisation of McBride acting on a deBruijn representation). The formalisation is here:<div><br></div><div>  <a href="https://wenkokke.github.io/sf/TypedFresh">https://wenkokke.github.io/sf/TypedFresh</a></div><div><br></div><div>Although more elaborate than I would like in parts, at least I succeed in proving progress and preservation for the system. Combining the two allows me to write a function that computes the reduction sequence for a term bounding by a given length. If the reduction sequence is longer than the specified length n, it returns the sequence of length n with the indicator "out of gas". I then tried to compute the reduction sequence for two plus two. </div><div><br></div><div>Computing the first four reduction steps takes a couple of minutes rather than a couple of seconds. Attempting to normalise the term fully (which should take fourteen reduction steps) does not complete even if run overnight!</div><div><br><pre class="gmail-Agda" style="margin:0px 0px 15px;padding:0.4em;font-size:0.85em;border:1px solid rgb(201,225,246);border-radius:3px;background:rgb(232,242,251);overflow-x:auto;font-family:"Source Code Pro",Consolas,Menlo,Monaco,"Lucida Console","Liberation Mono","DejaVu Sans Mono","Bitstream Vera Sans Mono","Courier New",monospace,serif;color:rgb(17,17,17);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;word-spacing:0px;text-decoration-style:initial;text-decoration-color:initial"><a id="gmail-19561" class="gmail-Comment" style="color:rgb(178,34,34);text-decoration:none">{-
_ : normalise 4 ⊢four ≡
  out-of-gas
  ((`Y
    (`λ 0 ⇒
     (`λ 1 ⇒ (`λ 2 ⇒ `if0 ` 1 then ` 2 else ` 0 · (`pred ` 1) · ` 2))))
   · (`suc (`suc `zero))
   · (`suc (`suc `zero))
   ⟶⟨ ξ-·₁ (ξ-·₁ (β-Y Fun refl)) ⟩
   (`λ 0 ⇒
    (`λ 1 ⇒
     `if0 ` 0 then ` 1 else
     (`Y
      (`λ 0 ⇒
       (`λ 1 ⇒ (`λ 2 ⇒ `if0 ` 1 then ` 2 else ` 0 · (`pred ` 1) · ` 2))))
     · (`pred ` 0)
     · ` 1))
   · (`suc (`suc `zero))
   · (`suc (`suc `zero))
   ⟶⟨ ξ-·₁ (β-⇒ (Suc (Suc Zero))) ⟩
   (`λ 0 ⇒
    `if0 `suc (`suc `zero) then ` 0 else
    (`Y
     (`λ 1 ⇒
      (`λ 2 ⇒ (`λ 3 ⇒ `if0 ` 2 then ` 3 else ` 1 · (`pred ` 2) · ` 3))))
    · (`pred (`suc (`suc `zero)))
    · ` 0)
   · (`suc (`suc `zero))
   ⟶⟨ β-⇒ (Suc (Suc Zero)) ⟩
   `if0 `suc (`suc `zero) then `suc (`suc `zero) else
   (`Y
    (`λ 0 ⇒
     (`λ 1 ⇒ (`λ 2 ⇒ `if0 ` 1 then ` 2 else ` 0 · (`pred ` 1) · ` 2))))
   · (`pred (`suc (`suc `zero)))
   · (`suc (`suc `zero))
   ⟶⟨ β-if0-suc (Suc Zero) ⟩
   (`Y
    (`λ 0 ⇒
     (`λ 1 ⇒ (`λ 2 ⇒ `if0 ` 1 then ` 2 else ` 0 · (`pred ` 1) · ` 2))))
   · (`pred (`suc (`suc `zero)))
   · (`suc (`suc `zero))
   ∎)
_ = refl
-}</a></pre><p style="margin:0px 0px 15px;padding:0px;color:rgb(17,17,17);font-family:-apple-system,system-ui,"Segoe UI",Roboto,Helvetica,Arial,sans-serif,"Apple Color Emoji","Segoe UI Emoji","Segoe UI Symbol";font-size:16px;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(253,253,253);text-decoration-style:initial;text-decoration-color:initial">To try to uncover the problem, I attempted a single step at a time. The first reduction step is fine.</p><pre class="gmail-Agda" style="margin:0px 0px 15px;padding:0.4em;font-size:0.85em;border:1px solid rgb(201,225,246);border-radius:3px;background:rgb(232,242,251);overflow-x:auto;font-family:"Source Code Pro",Consolas,Menlo,Monaco,"Lucida Console","Liberation Mono","DejaVu Sans Mono","Bitstream Vera Sans Mono","Courier New",monospace,serif;color:rgb(17,17,17);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;word-spacing:0px;text-decoration-style:initial;text-decoration-color:initial"><a id="gmail-20814" class="gmail-Comment" style="color:rgb(178,34,34);text-decoration:none">{-
_ : normalise 1 ⊢four ≡
  out-of-gas
  ((`Y
    (`λ 0 ⇒
     (`λ 1 ⇒ (`λ 2 ⇒ `if0 ` 1 then ` 2 else ` 0 · (`pred ` 1) · ` 2))))
   · (`suc (`suc `zero))
   · (`suc (`suc `zero))
   ⟶⟨ ξ-·₁ (ξ-·₁ (β-Y Fun refl)) ⟩
   (`λ 0 ⇒
    (`λ 1 ⇒
     `if0 ` 0 then ` 1 else
     (`Y
      (`λ 0 ⇒
       (`λ 1 ⇒ (`λ 2 ⇒ `if0 ` 1 then ` 2 else ` 0 · (`pred ` 1) · ` 2))))
     · (`pred ` 0)
     · ` 1))
   · (`suc (`suc `zero))
   · (`suc (`suc `zero))
   ∎)
  (⊢λ
   (⊢λ
    (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z)
     (⊢Y
      (⊢λ
       (⊢λ
        (⊢λ
         (⊢if0 (Ax (S m≢n Z)) (Ax Z)
          (Ax (S p≢n (S p≢m Z)) · ⊢pred (Ax (S m≢n Z)) · Ax Z)))))
      · ⊢pred (Ax (S (fresh-lemma CollectionDec.here) Z))
      · Ax Z)))
   · ⊢suc (⊢suc ⊢zero)
   · ⊢suc (⊢suc ⊢zero))
_ = refl
-}</a></pre>The second reduction step yields a huge term. It contains the free variable _x_1862 where I would have expected a fresh variable name to be computed. Looking at the four-step reduction above we see the name computed properly, but when going from the first step to the second it is not. The huge size of the type derivation suggests why it did not terminate after one day.<br><br>Help! Why doesn’t the term reduce? Is there some way to force it to reduce?</div><div><br></div><div>Cheers, -- P</div><div><br><pre class="gmail-Agda" style="margin:0px 0px 15px;padding:0.4em;font-size:0.85em;border:1px solid rgb(201,225,246);border-radius:3px;background:rgb(232,242,251);overflow-x:auto;font-family:"Source Code Pro",Consolas,Menlo,Monaco,"Lucida Console","Liberation Mono","DejaVu Sans Mono","Bitstream Vera Sans Mono","Courier New",monospace,serif;color:rgb(17,17,17);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;word-spacing:0px;text-decoration-style:initial;text-decoration-color:initial"><a id="gmail-22098" class="gmail-Comment" style="color:rgb(178,34,34);text-decoration:none">{-
_ : normalise 1
 (⊢λ
  (⊢λ
   (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z)
    (⊢Y
     (⊢λ
      (⊢λ
       (⊢λ
        (⊢if0 (Ax (S m≢n Z)) (Ax Z)
         (Ax (S p≢n (S p≢m Z)) · ⊢pred (Ax (S m≢n Z)) · Ax Z)))))
     · ⊢pred (Ax (S (fresh-lemma CollectionDec.here) Z))
     · Ax Z)))
  · ⊢suc (⊢suc ⊢zero)
  · ⊢suc (⊢suc ⊢zero))
  ≡
  out-of-gas
  ((`λ _x_1862 ⇒
    (`λ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) ⇒
     `if0 ` _x_1862 then
     ` (suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)) else
     (`Y
      (`λ 0 ⇒
       (`λ 1 ⇒ (`λ 2 ⇒ `if0 ` 1 then ` 2 else ` 0 · (`pred ` 1) · ` 2))))
     · (`pred ` _x_1862)
     · ` (suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005))))
   · (`suc (`suc `zero))
   · (`suc (`suc `zero))
   ⟶⟨ ξ-·₁ (β-⇒ (Suc (Suc Zero))) ⟩
   (`λ
    foldr _⊔_ 0
    (map suc
     (filter (λ x → ¬? (x ≟ _x_1862))
      (filter
       (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
       (_x_1862 ∷
        [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
        suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
       | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))))
    ⇒
    `if0
    (((λ w →
         ((λ x → ` x) , _x_1862 ↦ `suc (`suc `zero)) w | w ≟ _x_1862)
      , suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) ↦
      `
      foldr _⊔_ 0
      (map suc
       (filter (λ x → ¬? (x ≟ _x_1862))
        (filter
         (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
         (_x_1862 ∷
          [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
          suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
         | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
     _x_1862
     | _x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))
    then
    (((λ w →
         ((λ x → ` x) , _x_1862 ↦ `suc (`suc `zero)) w | w ≟ _x_1862)
      , suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) ↦
      `
      foldr _⊔_ 0
      (map suc
       (filter (λ x → ¬? (x ≟ _x_1862))
        (filter
         (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
         (_x_1862 ∷
          [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
          suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
         | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
     (suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))
     | suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) ≟
       suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))
    else
    (`Y
     (`λ
      suc
      (foldr _⊔_ 0
       (map suc
        (filter (λ x → ¬? (x ≟ _x_1862))
         (filter
          (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
          (_x_1862 ∷
           [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
           suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
          | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
      ⊔
      foldr _⊔_ 0
      (map suc
       (filter (λ x → ¬? (x ≟ _x_1862))
        (filter
         (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
         (_x_1862 ∷
          [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
          suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
         | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))))
      ⇒
      (`λ
       suc
       (suc
        (foldr _⊔_ 0
         (map suc
          (filter (λ x → ¬? (x ≟ _x_1862))
           (filter
            (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
            (_x_1862 ∷
             [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
             suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
            | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
        ⊔
        foldr _⊔_ 0
        (map suc
         (filter (λ x → ¬? (x ≟ _x_1862))
          (filter
           (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
           (_x_1862 ∷
            [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
            suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
           | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
       ⊔
       (suc
        (foldr _⊔_ 0
         (map suc
          (filter (λ x → ¬? (x ≟ _x_1862))
           (filter
            (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
            (_x_1862 ∷
             [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
             suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
            | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
        ⊔
        foldr _⊔_ 0
        (map suc
         (filter (λ x → ¬? (x ≟ _x_1862))
          (filter
           (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
           (_x_1862 ∷
            [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
            suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
           | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
       ⇒
       (`λ
        suc
        (suc
         (suc
          (foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
          ⊔
          foldr _⊔_ 0
          (map suc
           (filter (λ x → ¬? (x ≟ _x_1862))
            (filter
             (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
             (_x_1862 ∷
              [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
              suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
             | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
         ⊔
         (suc
          (foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
          ⊔
          foldr _⊔_ 0
          (map suc
           (filter (λ x → ¬? (x ≟ _x_1862))
            (filter
             (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
             (_x_1862 ∷
              [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
              suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
             | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))))))
        ⊔
        (suc
         (suc
          (foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
          ⊔
          foldr _⊔_ 0
          (map suc
           (filter (λ x → ¬? (x ≟ _x_1862))
            (filter
             (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
             (_x_1862 ∷
              [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
              suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
             | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
         ⊔
         (suc
          (foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
          ⊔
          foldr _⊔_ 0
          (map suc
           (filter (λ x → ¬? (x ≟ _x_1862))
            (filter
             (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
             (_x_1862 ∷
              [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
              suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
             | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))))))
        ⇒
        `if0
        `
        (suc
         (suc
          (foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
          ⊔
          foldr _⊔_ 0
          (map suc
           (filter (λ x → ¬? (x ≟ _x_1862))
            (filter
             (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
             (_x_1862 ∷
              [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
              suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
             | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
         ⊔
         (suc
          (foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
          ⊔
          foldr _⊔_ 0
          (map suc
           (filter (λ x → ¬? (x ≟ _x_1862))
            (filter
             (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
             (_x_1862 ∷
              [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
              suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
             | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))))))
        then
        `
        (suc
         (suc
          (suc
           (foldr _⊔_ 0
            (map suc
             (filter (λ x → ¬? (x ≟ _x_1862))
              (filter
               (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
               (_x_1862 ∷
                [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
                suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
               | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
           ⊔
           foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
          ⊔
          (suc
           (foldr _⊔_ 0
            (map suc
             (filter (λ x → ¬? (x ≟ _x_1862))
              (filter
               (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
               (_x_1862 ∷
                [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
                suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
               | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
           ⊔
           foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))))))
         ⊔
         (suc
          (suc
           (foldr _⊔_ 0
            (map suc
             (filter (λ x → ¬? (x ≟ _x_1862))
              (filter
               (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
               (_x_1862 ∷
                [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
                suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
               | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
           ⊔
           foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
          ⊔
          (suc
           (foldr _⊔_ 0
            (map suc
             (filter (λ x → ¬? (x ≟ _x_1862))
              (filter
               (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
               (_x_1862 ∷
                [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
                suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
               | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
           ⊔
           foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))))
        else
        `
        (suc
         (foldr _⊔_ 0
          (map suc
           (filter (λ x → ¬? (x ≟ _x_1862))
            (filter
             (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
             (_x_1862 ∷
              [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
              suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
             | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
         ⊔
         foldr _⊔_ 0
         (map suc
          (filter (λ x → ¬? (x ≟ _x_1862))
           (filter
            (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
            (_x_1862 ∷
             [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
             suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
            | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
        ·
        (`pred
         `
         (suc
          (suc
           (foldr _⊔_ 0
            (map suc
             (filter (λ x → ¬? (x ≟ _x_1862))
              (filter
               (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
               (_x_1862 ∷
                [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
                suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
               | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
           ⊔
           foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
          ⊔
          (suc
           (foldr _⊔_ 0
            (map suc
             (filter (λ x → ¬? (x ≟ _x_1862))
              (filter
               (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
               (_x_1862 ∷
                [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
                suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
               | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
           ⊔
           foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))))
        ·
        `
        (suc
         (suc
          (suc
           (foldr _⊔_ 0
            (map suc
             (filter (λ x → ¬? (x ≟ _x_1862))
              (filter
               (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
               (_x_1862 ∷
                [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
                suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
               | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
           ⊔
           foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
          ⊔
          (suc
           (foldr _⊔_ 0
            (map suc
             (filter (λ x → ¬? (x ≟ _x_1862))
              (filter
               (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
               (_x_1862 ∷
                [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
                suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
               | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
           ⊔
           foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))))))
         ⊔
         (suc
          (suc
           (foldr _⊔_ 0
            (map suc
             (filter (λ x → ¬? (x ≟ _x_1862))
              (filter
               (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
               (_x_1862 ∷
                [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
                suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
               | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
           ⊔
           foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
          ⊔
          (suc
           (foldr _⊔_ 0
            (map suc
             (filter (λ x → ¬? (x ≟ _x_1862))
              (filter
               (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
               (_x_1862 ∷
                [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
                suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
               | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
           ⊔
           foldr _⊔_ 0
           (map suc
            (filter (λ x → ¬? (x ≟ _x_1862))
             (filter
              (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
              (_x_1862 ∷
               [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
               suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
              | ¬?
                (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))))))))
    ·
    (`pred
     (((λ w →
          ((λ x → ` x) , _x_1862 ↦ `suc (`suc `zero)) w | w ≟ _x_1862)
       , suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) ↦
       `
       foldr _⊔_ 0
       (map suc
        (filter (λ x → ¬? (x ≟ _x_1862))
         (filter
          (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
          (_x_1862 ∷
           [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
           suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
          | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
      _x_1862
      | _x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
    ·
    (((λ w →
         ((λ x → ` x) , _x_1862 ↦ `suc (`suc `zero)) w | w ≟ _x_1862)
      , suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) ↦
      `
      foldr _⊔_ 0
      (map suc
       (filter (λ x → ¬? (x ≟ _x_1862))
        (filter
         (λ x → ¬? (x ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
         (_x_1862 ∷
          [ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) , _x_1862 ,
          suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ])
         | ¬? (_x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))))))
     (suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005))
     | suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_2005) ≟
       suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887)))
   · (`suc (`suc `zero))
   ∎)
  (⊢λ
   (⊢if0
    (.Typed.⊢ρ′
     (λ {w} w∈ {z₁} y∈ →
        .Typed.Σ
        (⊢λ
         (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016))
        (⊢suc (⊢suc ⊢zero)) w∈ y∈
        | w ≟ _x_1862)
     (.Typed.⊢ρ
      (⊢λ
       (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016))
      (⊢suc (⊢suc ⊢zero)))
     (λ {_} x∈ → x∈)
     (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016)
     (CollectionDec.\\-to-∷ ℕ _≟_ (λ {_} x∈ → x∈) CollectionDec.here
      | _x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))
     (S (fresh-lemma CollectionDec.here) Z)
     | _x_1862 ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))
    (.Typed.⊢ρ′
     (λ {w} w∈ {z₁} y∈ →
        .Typed.Σ
        (⊢λ
         (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016))
        (⊢suc (⊢suc ⊢zero)) w∈ y∈
        | w ≟ _x_1862)
     (.Typed.⊢ρ
      (⊢λ
       (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016))
      (⊢suc (⊢suc ⊢zero)))
     (λ {_} x∈ → x∈)
     (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016)
     (CollectionDec.\\-to-∷ ℕ _≟_ (λ {_} x∈ → x∈)
      (CollectionDec.there CollectionDec.here)
      | suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) ≟
        suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))
     Z
     | suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887) ≟
       suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))
    (⊢subst
     (λ {w} w∈′ {z₁} →
        .Typed.Σ′
        (λ {w₁} w∈ {z₂} y∈ →
           .Typed.Σ
           (⊢λ
            (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016))
           (⊢suc (⊢suc ⊢zero)) w∈ y∈
           | w₁ ≟ _x_1862)
        (.Typed.⊢ρ
         (⊢λ
          (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016))
         (⊢suc (⊢suc ⊢zero)))
        (λ {_} x∈ → x∈)
        (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016) w∈′
        | w ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))
     (.Typed.⊢ρ′
      (λ {w} w∈ {z₁} y∈ →
         .Typed.Σ
         (⊢λ
          (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016))
         (⊢suc (⊢suc ⊢zero)) w∈ y∈
         | w ≟ _x_1862)
      (.Typed.⊢ρ
       (⊢λ
        (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016))
       (⊢suc (⊢suc ⊢zero)))
      (λ {_} x∈ → x∈)
      (⊢if0 (Ax (S (fresh-lemma CollectionDec.here) Z)) (Ax Z) _2016))
     (λ {z₁} x →
        CollectionDec.\\-to-∷ ℕ _≟_ (λ {_} x∈ → x∈)
        (CollectionDec.there (CollectionDec.there x))
        | z₁ ≟ suc _x_1862 ⊔ foldr _⊔_ 0 (map suc _xs_1887))
     _2016))
   · ⊢suc (⊢suc ⊢zero))
-}</a></pre><br></div><div><div><br><br>.   \ Philip Wadler, Professor of Theoretical Computer Science,<br>.   /\ School of Informatics, University of Edinburgh<br>.  /  \ and Senior Research Fellow, IOHK<div><div><div class="gmail_signature"><div dir="ltr"><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>
</div></div></div></div>