|
1 | 1 | import Base |
| 2 | +import Mathlib.Data.Nat.Factorial.Basic |
2 | 3 |
|
3 | 4 | /-! |
4 | 5 |
|
@@ -26,3 +27,170 @@ divergent def factorial (n : U64) : Result U64 := |
26 | 27 | n * i1 |
27 | 28 |
|
28 | 29 | end factorial |
| 30 | + |
| 31 | +/-! |
| 32 | +
|
| 33 | +We will need some lemmas on factorials lifted to ℤ and U64 |
| 34 | +to prove theorems about factorial. |
| 35 | +
|
| 36 | +This will be useful to introduce arithmetic work in Aeneas. |
| 37 | +
|
| 38 | +First, we would like to prove that the factorial of natural number as integers possess |
| 39 | +the recursive property of factorials. |
| 40 | +-/ |
| 41 | + |
| 42 | +lemma Int.mul_factorial_pred: ∀ {n: ℤ}, 0 < n → n * (n - 1).toNat.factorial = n.toNat.factorial := by |
| 43 | + intro n n_pos |
| 44 | + -- This is a generic tactic to lift types to another types while discharging conditions, e.g. positivity, etc. |
| 45 | + lift (n: ℤ) to ℕ |
| 46 | + -- `omega` is another generic tactic to close arithmetic goals automatically, `linarith` is another option |
| 47 | + -- for linear arithmetic. |
| 48 | + omega |
| 49 | + -- `simp` here is used as a normalization tactic, then we normalize the casts themselves |
| 50 | + -- and finally, we use the natural number version of the recursive property. |
| 51 | + -- `exact_mod_cast` is a "here's the term modulo casts" tactic. |
| 52 | + simp; norm_cast; exact Nat.mul_factorial_pred (by exact_mod_cast n_pos) |
| 53 | + |
| 54 | + |
| 55 | +lemma U64.mul_factorial_pred: ∀ {n: U64}, 0#u64 < n → (n: ℤ) * (n.toNat - 1).factorial = n.toNat.factorial := by |
| 56 | + intro n n_pos |
| 57 | + -- `convert` is a "here's the term, please match it as much as you can and let me discharge the differences". |
| 58 | + convert (Int.mul_factorial_pred _) |
| 59 | + . simp |
| 60 | + . scalar_tac |
| 61 | + |
| 62 | +/-! |
| 63 | +
|
| 64 | +We will need a mathematical theorem relating machine integers |
| 65 | +as Rust cannot compute factorials larger than 2^64 - 1 using u64s. |
| 66 | +
|
| 67 | +As factorial is a nice function, i.e. is increasing, we can just reduce |
| 68 | +the proof to proving that our domain's max value fits in a U64, i.e. 20! ≤ 2^64 - 1. |
| 69 | +
|
| 70 | +In Lean, certain goals can be "decided", e.g. we can just use a fast algorithm |
| 71 | +to produce a proof certificate that something is true. |
| 72 | +-/ |
| 73 | +theorem factorial_bounds: ∀ {n: ℕ}, n ≤ 20 → n.factorial ≤ U64.max := by |
| 74 | + intro n n_pos |
| 75 | + -- Our strategy will be to write this in ℕ to reuse natural number theorems. |
| 76 | + rw [U64.max] |
| 77 | + norm_cast |
| 78 | + -- We use transitivity of ≤ to use the natural number theorem _and_ decision. |
| 79 | + transitivity |
| 80 | + . exact Nat.factorial_le n_pos |
| 81 | + . decide |
| 82 | + |
| 83 | +/-! |
| 84 | +
|
| 85 | +We define a restricted scalar factorial to U64. |
| 86 | +To construct a scalar, we need to prove that it fits both of the ends of the scalar. |
| 87 | +For unsigned scalar, we only need to prove positivity _and_ it fits the max. |
| 88 | +
|
| 89 | +We leave the problem to `scalar_tac`, an Aeneas tactic that performs pre-processing step |
| 90 | +and knows facts about scalars to discharge scalar-related goals. This tactic calls internally `omega`, |
| 91 | +so it is _at least_ powerful as `omega`. |
| 92 | +
|
| 93 | +For the upper-bound, we reuse our just-proven theorem. |
| 94 | +-/ |
| 95 | +def scalar_factorial {n: ℕ} (bounds: n ≤ 20): U64 := |
| 96 | + Scalar.ofIntCore n.factorial ⟨ by scalar_tac, factorial_bounds bounds ⟩ |
| 97 | + |
| 98 | +/-! |
| 99 | +
|
| 100 | +We are ready to create the specification of factorial. |
| 101 | +
|
| 102 | +That is, under the precondition of having an integer smaller than 20, the Rust's extracted factorial function |
| 103 | +coincides with the theoretical scalar factorial function lifted from the mathematical scalar factorial function. |
| 104 | +
|
| 105 | +Nonetheless, we will not be able to finish the formalization with the proposed signature. |
| 106 | +
|
| 107 | +`@[pspec]`, a decorator, marks the following theorem as a specification, which can be searched by Aeneas tactics on specifications. |
| 108 | +-/ |
| 109 | + |
| 110 | +@[pspec] |
| 111 | +theorem factorial_spec_mistaken: ∀ (n: ℕ), |
| 112 | + (bounds: n ≤ 20) → |
| 113 | + -- We cast the input natural number it in a scalar. |
| 114 | + factorial.factorial (Scalar.ofIntCore n ⟨ by scalar_tac, by scalar_tac ⟩) = Result.ok (scalar_factorial bounds) := by |
| 115 | + intro n bounds |
| 116 | + -- We unfold the Rust definition. |
| 117 | + rw [factorial.factorial] |
| 118 | + -- Distinguish the trivial case and the complicated case. |
| 119 | + by_cases hzero : n = 0 |
| 120 | + . |
| 121 | + -- We let Lean reduce this to a trivial equation. |
| 122 | + simp [hzero, scalar_factorial] |
| 123 | + . |
| 124 | + -- We simplify the branching using our negation hypothesis. |
| 125 | + simp [hzero] |
| 126 | + /- |
| 127 | + Here, we need to reuse a specification on scalar substraction regarding `n - 1` |
| 128 | + `k` will be the new integer in the context and `Hk` the term that links `k` |
| 129 | + to its value. |
| 130 | + This specification can be discovered automatically and pasted by `progress?`. |
| 131 | + `progress` is one of the Aeneas tactics on specifications. |
| 132 | + -/ |
| 133 | + progress with Primitives.U64.sub_spec as ⟨ k, Hk ⟩ |
| 134 | + -- We have a problem here, the goal ask us to reuse our specification |
| 135 | + -- but with a `k`, not with a `Scalar.ofIntCore ...` |
| 136 | + -- rw [Hk] |
| 137 | + -- progress with factorial_spec |
| 138 | + stop |
| 139 | + |
| 140 | +/-! |
| 141 | +In general, specifications in Aeneas are written with an existential-style, that is, we will write there's a result |
| 142 | +to some operation and we will attach additional conditions to the value of the result. |
| 143 | +
|
| 144 | +This avoids to depend on things like proofs of boundness of a given scalar which could be arbitrary. |
| 145 | +
|
| 146 | +Here's a fixed proposal. |
| 147 | +-/ |
| 148 | + |
| 149 | +@[pspec] |
| 150 | +theorem factorial_spec: ∀ (n: U64), (bounds: n.val ≤ 20) → ∃ v, factorial.factorial n = Result.ok v ∧ v.val = n.val.toNat.factorial := by |
| 151 | + intro n bounds |
| 152 | + rw [factorial.factorial] |
| 153 | + by_cases hzero : n = 0#u64 |
| 154 | + . |
| 155 | + -- We let Lean reduce this to a trivial equation. |
| 156 | + simp [hzero, scalar_factorial] |
| 157 | + . |
| 158 | + -- We simplify the branching using our negation hypothesis. |
| 159 | + simp [hzero] |
| 160 | + -- Here, we need to reuse a specification on scalar substraction regarding `n - 1` |
| 161 | + -- `k` will be the new integer in the context and `Hk` the term that links `k` |
| 162 | + -- to its value. |
| 163 | + -- This specification can be discovered automatically and pasted by `progress?`. |
| 164 | + progress with Primitives.U64.sub_spec as ⟨ k, Hk ⟩ |
| 165 | + -- Here, we have the recursive call to `factorial` |
| 166 | + -- Lean will figure out automatically that calling ourselves with a smaller argument |
| 167 | + -- is sufficient to prove termination of this specification proof. |
| 168 | + progress with factorial_spec as ⟨ p, Hp ⟩ |
| 169 | + -- We still need to prove that np is a valid U64 scalar |
| 170 | + -- and its value is what is expected. |
| 171 | + progress with Primitives.U64.mul_spec as ⟨ y, Hy ⟩ |
| 172 | + -- We replace our collected equalities in the previous progresses. |
| 173 | + . |
| 174 | + rw [Hp, Hk] |
| 175 | + -- We do various wrangling with the equalities to simplify |
| 176 | + -- their casts, their form, remove the definitions (U64.max). |
| 177 | + simp only [Scalar.ofInt_val_eq, Int.pred_toNat] |
| 178 | + -- We leave it to Aeneas automation to discharge the non-negativity of the scalar. |
| 179 | + rw [U64.mul_factorial_pred (by scalar_tac)] |
| 180 | + simp only [U64.toNat] |
| 181 | + -- Our original bound live in ℤ, we need it in ℕ, let's lift it quickly. |
| 182 | + have bounds' : n.toNat ≤ 20 := by { |
| 183 | + -- Let's lift the goal, i.e. the inequality in ℕ, to ℤ and conclude immediately. |
| 184 | + zify; simp [bounds] |
| 185 | + } |
| 186 | + exact factorial_bounds bounds' |
| 187 | + . rw [Hy, Hp, Hk] |
| 188 | + simp only [Scalar.ofInt_val_eq, Int.pred_toNat] |
| 189 | + rw [U64.mul_factorial_pred (by scalar_tac)] |
| 190 | +-- We need to prove that we are calling the specification with a smaller scalar. |
| 191 | +-- Let's prove it over the cast to natural numbers. |
| 192 | +termination_by n => n.val.toNat |
| 193 | +decreasing_by |
| 194 | + simp_wf -- We simplify the well-founded relation definition. |
| 195 | + rw [Hk] -- Inject the definition of `k`. |
| 196 | + simp -- Let a Lean tactic conclude trivially. |
0 commit comments