├── .gitignore ├── Calcify.lean ├── Calcify ├── Demo.lean ├── Eggsamples.lean └── Tests.lean ├── LICENSE ├── README.md ├── lake-manifest.json ├── lakefile.toml └── lean-toolchain /.gitignore: -------------------------------------------------------------------------------- 1 | /.lake 2 | -------------------------------------------------------------------------------- /Calcify.lean: -------------------------------------------------------------------------------- 1 | import Lean.Meta.Tactic.TryThis 2 | import Lean.Elab.Tactic.ShowTerm 3 | import Lean.Elab.Tactic.Guard 4 | import Lean 5 | 6 | open Lean Elab Tactic Meta 7 | 8 | open Lean.Meta.Tactic.TryThis (delabToRefinableSyntax addSuggestion) 9 | 10 | 11 | -- copied from std4/Std/Lean/Meta/Basic.lean 12 | /-- Solve a goal by synthesizing an instance. -/ 13 | -- FIXME: probably can just be `g.inferInstance` once leanprover/lean4#2054 is fixed 14 | private def Lean.MVarId.synthInstance (g : MVarId) : MetaM Unit := do 15 | g.assign (← Lean.Meta.synthInstance (← g.getType)) 16 | 17 | -- NB: Pattern matching on terms using `mkAppN` is not good practice, as 18 | -- it generates very large and inefficient code. 19 | 20 | partial def mkEqTrans' (p₁ p₂ : Expr) : MetaM Expr := do 21 | match_expr p₁ with 22 | | Eq.trans _ _ _ _ p₁₁ p₁₂ => mkEqTrans' p₁₁ (← mkEqTrans' p₁₂ p₂) 23 | | _ => mkEqTrans p₁ p₂ 24 | 25 | partial def mkHEqTrans' (p₁ p₂ : Expr) : MetaM Expr := do 26 | match_expr p₁ with 27 | | HEq.trans _ _ _ _ p₁₁ p₁₂ => mkHEqTrans' p₁₁ (← mkHEqTrans' p₁₂ p₂) 28 | | _ => mkHEqTrans p₁ p₂ 29 | 30 | partial def mkCongrArg' (f p : Expr) : MetaM Expr := do 31 | -- The function is constant? This becomes refl 32 | if let .lam _ _ b _ := f then 33 | if ! b.hasLooseBVars then 34 | return ← mkEqRefl b 35 | 36 | -- The function is the identity? Short-circuit 37 | if let .lam _ _ (.bvar 0) _ := f then 38 | return p 39 | 40 | -- Push congr into transitivity 41 | match_expr p with 42 | | Eq.trans _ _ _ _ p₁ p₂ => do 43 | return ← mkEqTrans' (← mkCongrArg' f p₁) (← mkCongrArg' f p₂) 44 | | _ => mkCongrArg f p 45 | 46 | -- congrFun is a special case of congrArg 47 | def mkCongrFun' (h x : Expr) : MetaM Expr := do 48 | let some (α, _f₁, _f₂) := (← inferType h).eq? | throwError "Expected proof of equality" 49 | mkCongrArg' (.lam `f α (.app (.bvar 0) x) .default) h 50 | 51 | -- congr can be written as a composition of congrFun and congrArg 52 | def mkCongr' (x₁ f₂ : Expr) (p1 p2 : Expr) : MetaM Expr := do 53 | mkEqTrans' (← mkCongrFun' p1 x₁) (← mkCongrArg' f₂ p2) 54 | 55 | 56 | def mkEqOfHEq' (h : Expr) : MetaM Expr := do 57 | match_expr h with 58 | | HEq.refl _ x => mkEqRefl x 59 | | heq_of_eq _ _ _ h => pure h 60 | | _ => mkEqOfHEq h 61 | 62 | def mkHEqOfEq' (h : Expr) : MetaM Expr := do 63 | match_expr h with 64 | | Eq.refl _ x => mkHEqRefl x 65 | | eq_of_heq _ _ _ h => pure h 66 | | _ => mkAppM ``heq_of_eq #[h] 67 | 68 | def mkFunExt' (p : Expr) : MetaM Expr := do 69 | if let .lam n t (mkApp6 (.const ``Eq.trans _) _ _ _ _ p1 p2) bi := p then 70 | return ← mkEqTrans' 71 | (← mkFunExt' (.lam n t p1 bi)) 72 | (← mkFunExt' (.lam n t p2 bi)) 73 | mkFunExt p 74 | 75 | partial def mkEqSymm' (h : Expr) : MetaM Expr := do 76 | match_expr h with 77 | | Eq.symm _ _ _ h => pure h 78 | | Eq.trans _ _ _ _ p₁ p₂ => mkEqTrans' (← mkEqSymm' p₂) (← mkEqSymm' p₁) 79 | | congrArg _ _ _ _ f p1 => mkCongrArg' f (← mkEqSymm' p1) 80 | | _ => mkEqSymm h 81 | 82 | def mkPropExt' (e : Expr) : MetaM Expr := do 83 | match_expr e with 84 | | Iff.refl p => mkEqRefl p 85 | | Iff.of_eq _ _ p => pure p 86 | | _ => mkPropExt e 87 | 88 | partial def mkEqMPR' (e1 e2 : Expr) : MetaM Expr := do 89 | match_expr e1 with 90 | | congrArg _ _ _ _ f p1 => do 91 | -- A mpr applied to an congruence with equality can be turned into transitivities 92 | if let .lam n t (mkApp3 (.const ``Eq _) _β b₁ b₂) bi := f then 93 | return ← mkEqTrans' (← mkCongrArg' (.lam n t b₁ bi) p1) 94 | (← mkEqTrans' e2 95 | (← mkCongrArg' (.lam n t b₂ bi) (← mkEqSymm' p1))) 96 | -- same with HEq 97 | if let .lam n t (mkApp4 (.const ``HEq _) _β₁ b₁ _β₂ b₂) bi := f then 98 | return ← mkHEqTrans' (← mkHEqOfEq' (← mkCongrArg' (.lam n t b₁ bi) p1)) 99 | (← mkHEqTrans' e2 100 | (← mkHEqOfEq' (← mkCongrArg' (.lam n t b₂ bi) (← mkEqSymm' p1)))) 101 | -- same with Iff 102 | if let .lam n t (mkApp2 (.const ``Iff _) b₁ b₂) bi := f then 103 | return ← mkIffOfEq ( 104 | ← mkEqTrans' (← mkCongrArg' (.lam n t b₁ bi) p1) 105 | (← mkEqTrans' (← mkPropExt' e2) 106 | (← mkCongrArg' (.lam n t b₂ bi) (← mkEqSymm' p1)))) 107 | 108 | -- Special case of the above, with an eta-contracted congruence 109 | match_expr f with 110 | | Eq _β _b₁ => return ← mkEqTrans' e2 (← mkEqSymm' p1) 111 | | Iff _ => return ← mkIffOfEq (← mkEqTrans' (← mkPropExt' e2) (← mkEqSymm' p1)) 112 | | _ => pure () 113 | | _ => pure () 114 | 115 | -- A mpr applied to an mpr can be turned into a transitivity 116 | match_expr e2 with 117 | | Eq.mpr _ _ p2 p3 => return ← mkEqMPR' (← mkEqTrans' e1 p2) p3 118 | | _ => pure () 119 | 120 | mkEqMPR e1 e2 121 | 122 | def mkEqNDRec' (motive h1 h2 : Expr) : MetaM Expr := do 123 | -- TODO: Eq.mpr (congrArg …) is just Eq.ndrec, is it? 124 | -- So maybe the mkEqMPR handling above should be moved here 125 | mkEqMPR' (← mkCongrArg motive (← mkEqSymm' h2)) h1 126 | 127 | /- 128 | The following four functions push `ite_congr` past transitivity. 129 | It is quite ugly and hairy, and should be generalized to arbitrary `congr` 130 | lemmas. 131 | -/ 132 | 133 | partial def mkIteCongr1 (goal : MVarId) (p : Expr) : MetaM Unit := do 134 | match_expr p with 135 | | Eq.refl _ _ => goal.refl 136 | | Eq.trans _ _ _ _ p1 p2 => do 137 | let [goal1, goal2, _] ← goal.applyConst `Eq.trans | throwError "Could not apply trans" 138 | mkIteCongr1 goal1 p1 139 | mkIteCongr1 goal2 p2 140 | | _ => do 141 | -- logInfo m!"here at {p}\n{goal}" 142 | let e ← mkConstWithFreshMVarLevels ``ite_congr 143 | let (mvars, _) ← forallMetaTelescopeReducing (← inferType e) 144 | assert! mvars.size == 12 145 | let _ ← isDefEq (mkMVar goal) (mkAppN e mvars) 146 | let _ ← isDefEq mvars[9]! p 147 | mvars[8]!.mvarId!.synthInstance 148 | (← mvars[10]!.mvarId!.intro1).2.refl 149 | (← mvars[11]!.mvarId!.intro1).2.refl 150 | -- logInfo m!"{mvars}" 151 | 152 | partial def mkIteCongr2 (goal : MVarId) (p : Expr) : MetaM Unit := do 153 | if let .lam n t b bi := p then 154 | match_expr b with 155 | | Eq.refl _ _ => do 156 | goal.refl 157 | return 158 | | Eq.trans _ _ _ _ p1 p2 => do 159 | let [goal1, goal2, _] ← goal.applyConst `Eq.trans | throwError "Could not apply trans" 160 | mkIteCongr2 goal1 (.lam n t p1 bi) 161 | mkIteCongr2 goal2 (.lam n t p2 bi) 162 | return 163 | | _ => 164 | unless b.hasLooseBVars do 165 | -- Rewriting isn't actually contextual, use simple congr 166 | let some (_, lhs, _rhs) := (← goal.getType).eq? 167 | | throwError "Expected equality goal, got{indentExpr (← goal.getType)}" 168 | match_expr lhs with 169 | | ite α c inst _t e => 170 | let u ← getLevel α 171 | let _ ← isDefEq (mkMVar goal) 172 | (← mkCongrFun' (← mkCongrArg' (mkApp3 (.const ``ite [u]) α c inst) b) e) 173 | | _ => pure () 174 | pure () 175 | 176 | let e ← mkConstWithFreshMVarLevels ``ite_congr 177 | let (mvars, _) ← forallMetaTelescopeReducing (← inferType e) 178 | assert! mvars.size == 12 179 | let _ ← isDefEq (mkMVar goal) (mkAppN e mvars) 180 | mvars[9]!.mvarId!.refl 181 | mvars[8]!.mvarId!.synthInstance 182 | _ ← isDefEq mvars[10]! p 183 | (← mvars[11]!.mvarId!.intro1).2.refl 184 | 185 | partial def mkIteCongr3 (goal : MVarId) (p : Expr) : MetaM Unit := do 186 | if let .lam n t b bi := p then 187 | match_expr b with 188 | | Eq.refl _ _ => do 189 | goal.refl 190 | return 191 | | Eq.trans _ _ _ _ p1 p2 => do 192 | let [goal1, goal2, _] ← goal.applyConst `Eq.trans | throwError "Could not apply trans" 193 | mkIteCongr3 goal1 (.lam n t p1 bi) 194 | mkIteCongr3 goal2 (.lam n t p2 bi) 195 | return 196 | | _ => 197 | unless b.hasLooseBVars do 198 | -- Rewriting isn't actually contextual, use simple congr 199 | let some (_, lhs, _rhs) := (← goal.getType).eq? 200 | | throwError "Expected equality goal, got{indentExpr (← goal.getType)}" 201 | match_expr lhs with 202 | | ite α c inst t _e => 203 | let u ← getLevel α 204 | let _ ← isDefEq (mkMVar goal) 205 | (← mkCongrArg' (mkApp4 (.const ``ite [u]) α c inst t) b) 206 | | _ => pure () 207 | 208 | let e ← mkConstWithFreshMVarLevels ``ite_congr 209 | let (mvars, _) ← forallMetaTelescopeReducing (← inferType e) 210 | assert! mvars.size == 12 211 | let _ ← isDefEq (mkMVar goal) (mkAppN e mvars) 212 | mvars[9]!.mvarId!.refl 213 | mvars[8]!.mvarId!.synthInstance 214 | (← mvars[10]!.mvarId!.intro1).2.refl 215 | _ ← isDefEq mvars[11]! p 216 | 217 | def mkIteCongr (goal p1 p2 p3 : Expr) : MetaM Expr := do 218 | let mvar ← mkFreshExprSyntheticOpaqueMVar goal 219 | let [goal1, goal2, _] ← mvar.mvarId!.applyConst `Eq.trans | throwError "Could not apply trans" 220 | mkIteCongr1 goal1 p1 221 | let [goal2, goal3, _] ← goal2.applyConst `Eq.trans| throwError "Could not apply trans" 222 | mkIteCongr2 goal2 p2 223 | mkIteCongr3 goal3 p3 224 | instantiateMVars mvar 225 | 226 | partial def mkOfEqTrue' (p : Expr) : MetaM Expr := do 227 | match_expr p with 228 | | eq_self _α a => mkEqRefl a 229 | | iff_self a => pure <| mkApp (mkConst ``Iff.refl) a 230 | | eq_true _P p => pure p 231 | | Eq.trans _ _ _ _ p1 p2 => do mkEqMPR' p1 (← mkOfEqTrue' p2) 232 | | _ => do mkOfEqTrue p 233 | 234 | partial def simplify (e : Expr) : MetaM Expr := do 235 | lambdaTelescope e fun xs e => do 236 | let e := e.headBeta 237 | let e' ← match_expr e with 238 | 239 | -- eliminate id application, and hope for the best 240 | | id _ p => simplify p 241 | 242 | -- Use the smart constructors above 243 | | congr _α _β _f₁ f₂ x₁ _x₂ p1 p2 => do mkCongr' x₁ f₂ (← simplify p1) (← simplify p2) 244 | | of_eq_true _ p => do mkOfEqTrue' (← simplify p) 245 | | congrFun _ _ _ _ p1 x => do mkCongrFun' (← simplify p1) x 246 | | congrArg _α _β _a _a' f p => do mkCongrArg' f (← simplify p) 247 | | funext _ _ _ _ p => do mkFunExt' (← simplify p) 248 | | Eq.mpr _ _ p₁ p₂ => do mkEqMPR' (← simplify p₁) (← simplify p₂) 249 | | Eq.refl _ _ => pure e 250 | | Eq.symm _ _ _ h => do mkEqSymm' (← simplify h) 251 | | Eq.trans _α _a _b _c p1 p2 => do mkEqTrans' (← simplify p1) (← simplify p2) 252 | | ite_congr _α _b _c _x _y _u _v _i1 _i2 p1 p2 p3 => 253 | mkIteCongr (← inferType e) (← simplify p1) (← simplify p2) (← simplify p3) 254 | | HEq.refl _ _ => pure e 255 | | HEq.trans _α _β _γ _a _b _c p1 p2 => do mkHEqTrans' (← simplify p1) (← simplify p2) 256 | | eq_of_heq _α _a _b h => do mkEqOfHEq' (← simplify h) 257 | | heq_of_eq _α _a _b h => do mkHEqOfEq' (← simplify h) 258 | | _ => 259 | -- This can have extra arguments 260 | if e.isAppOf ``Eq.ndrec && e.getAppNumArgs ≥ 6 then 261 | let xs := e.getAppArgs 262 | let motive := xs[2]! 263 | let m := xs[3]! 264 | let h ← simplify xs[5]! 265 | if h.isAppOf ``Eq.refl then 266 | return ← simplify (mkAppN m xs[6:]) 267 | 268 | -- beta-reduce through Eq.ndrec 269 | -- (TODO: Could do more arguments in one go) 270 | if e.getAppNumArgs > 6 then 271 | let arg := xs[6]! 272 | if let .lam n d motiveType bi := motive then 273 | if motiveType.isForall && !motiveType.bindingDomain!.hasLooseBVars then 274 | let motive' := .lam n d (motiveType.bindingBody!.instantiate1 arg) bi 275 | let m' := m.beta #[arg] 276 | let e' := mkAppN (← mkEqNDRec motive' m' h) xs[7:] 277 | return ← simplify e' 278 | 279 | return ← simplify (mkAppN (← mkEqNDRec' motive m h) xs[6:]) 280 | 281 | -- Let's look through auxLemmas which are created by some tactics 282 | if let some e' ← delta? e (· matches .num (.str _ "_auxLemma") _) then 283 | return ← simplify e' 284 | 285 | -- unless e.getAppFn.isFVar do logInfo m!"Unrecognized: {e}" 286 | trace[calcify] "unrecognized:{indentExpr e}" 287 | pure e 288 | unless e == e' do 289 | trace[calcify] "simplify:{indentExpr e}\n==>{indentExpr e'}" 290 | mkLambdaFVars xs e' 291 | 292 | open Lean.Parser.Tactic 293 | 294 | partial def getCalcProof (proof : Expr) : MetaM Term := 295 | match_expr proof with 296 | | Eq.symm _ _ _ h => do 297 | let h ← getCalcProof h 298 | `($(h).$(mkIdent `symm)) 299 | /- 300 | Too bad we don't have congrArg in the Eq namespace 301 | | congrArg _ _ _ _ f h => do 302 | let h ← getCalcProof h 303 | let f ← delabToRefinableSyntax f 304 | `($(h).$(mkIdent `congrArg) $f) 305 | -/ 306 | | _ => delabToRefinableSyntax proof 307 | 308 | partial def getCalcSteps (proof : Expr) (acc : Array (TSyntax ``calcStep)) : 309 | MetaM (Array (TSyntax ``calcStep)) := 310 | match_expr proof with 311 | | Eq.trans _ _ rhs _ proof p2 => do 312 | let step ← `(calcStep|_ = $(← delabToRefinableSyntax rhs) := $(← getCalcProof proof)) 313 | getCalcSteps p2 (acc.push step) 314 | | _ => do 315 | let type ← whnf (← inferType proof) 316 | let some (_, _, rhs) := type.eq? | throwError "Expected proof of equality, got {type}" 317 | let step ← `(calcStep|_ = $(← delabToRefinableSyntax rhs) := $(← getCalcProof proof)) 318 | return acc.push step 319 | 320 | def delabCalcProof (e : Expr) : MetaM (TSyntax `tactic) := do 321 | let type ← whnf (← inferType e) 322 | let some (_, lhs, _) := type.eq? | throwError "Expected proof of equality, got {type}" 323 | let stepStx ← getCalcSteps e #[] 324 | `(tactic|calc 325 | $(← delabToRefinableSyntax lhs):term 326 | $stepStx*) 327 | 328 | def delabOfIffOfEqCalcProof (e : Expr) : MetaM (TSyntax `tactic) := do 329 | let type ← whnf (← inferType e) 330 | let some (_, lhs, rhs) := type.eq? | throwError "Expected proof of equality, got {type}" 331 | let stepStx ← getCalcSteps e #[] 332 | let finalStep ← `(calcStep|_ ↔ $(← delabToRefinableSyntax rhs) := $(mkIdent ``Iff.rfl)) 333 | let stepStx := stepStx.push finalStep 334 | `(tactic|calc 335 | $(← delabToRefinableSyntax lhs):term 336 | $stepStx*) 337 | 338 | def delabCalcTerm (e : Expr) : MetaM (TSyntax `term) := do 339 | let type ← whnf (← inferType e) 340 | let some (_, lhs, _) := type.eq? | throwError "Expected proof of equality, got {type}" 341 | let stepStx ← getCalcSteps e #[] 342 | `(term|calc 343 | $(← delabToRefinableSyntax lhs):term 344 | $stepStx*) 345 | 346 | def delabMPRCalc (p1 p2 : Expr) : MetaM (TSyntax ``tacticSeq) := do 347 | let t ← delabCalcProof p1 348 | if p2.isMVar then 349 | `(tacticSeq|conv => tactic => $t:tactic) 350 | else 351 | let restProof ← delabToRefinableSyntax p2 352 | `(tacticSeq|conv => tactic => $t:tactic 353 | refine $restProof) 354 | 355 | def delabOfEqTrueCalc (p : Expr) : MetaM (TSyntax ``tacticSeq) := do 356 | let t ← delabCalcProof p 357 | `(tacticSeq|apply $(mkIdent ``of_eq_true) 358 | $t) 359 | 360 | def delabCalcSeq (p : Expr) : MetaM (TSyntax ``tacticSeq) := do 361 | let t ← delabCalcProof p 362 | `(tacticSeq|$t:tactic) 363 | 364 | def delabTrivial (p : Expr) : MetaM (TSyntax ``tacticSeq) := do 365 | `(tacticSeq|exact $(← delabToRefinableSyntax p)) 366 | 367 | def delabProof (e : Expr) : MetaM (TSyntax ``tacticSeq) := do 368 | match_expr e with 369 | | True.intro => delabTrivial e 370 | | Iff.refl _ => delabTrivial e 371 | | Eq.mpr _ _ p1 p2 => delabMPRCalc p1 p2 372 | | Iff.of_eq _ _ p => 373 | let t ← delabOfIffOfEqCalcProof p 374 | `(tacticSeq|$t:tactic) 375 | | of_eq_true h p => 376 | if h.isEq then 377 | delabCalcSeq e 378 | else 379 | delabOfEqTrueCalc p 380 | | _ => do 381 | delabCalcSeq e 382 | 383 | elab (name := calcifyTac) tk:"calcify " t:tacticSeq : tactic => withMainContext do 384 | let goalMVar ← getMainGoal 385 | evalTactic t 386 | let proof ← instantiateMVars (mkMVar goalMVar) 387 | let proof ← simplify proof 388 | check proof 389 | let tactic ← delabProof proof 390 | 391 | /- 392 | let goal ← instantiateMVars (← goalMVar.getType) 393 | let testMVar ← mkFreshExprSyntheticOpaqueMVar goal 394 | withRef tk do 395 | Lean.Elab.Term.runTactic testMVar.mvarId! (← `(term|by $tactic)) 396 | -/ 397 | 398 | addSuggestion tk tactic (origSpan? := ← getRef) 399 | 400 | initialize registerTraceClass `calcify 401 | -------------------------------------------------------------------------------- /Calcify/Demo.lean: -------------------------------------------------------------------------------- 1 | import Calcify 2 | 3 | /-- 4 | info: Try this: calc 5 | 0 + n 6 | _ = n := (Nat.zero_add n) 7 | _ = 1 * n := (Nat.one_mul n).symm 8 | _ = 1 * 1 * n := congrArg (fun x => x * n) (Eq.symm (Nat.mul_one 1)) 9 | -/ 10 | #guard_msgs in 11 | example (n : Nat) : 0 + n = 1 * 1 * n := by 12 | calcify simp 13 | 14 | 15 | /-- 16 | info: Try this: calc 17 | 0 + n 18 | _ = n := (Nat.zero_add n) 19 | _ = 1 * n := (Nat.one_mul n).symm 20 | _ = 1 * 1 * n := congrArg (fun x => x * n) (Eq.symm (Nat.mul_one 1)) 21 | -/ 22 | #guard_msgs in 23 | example (n : Nat) : 0 + n = 1 * 1 * n := by 24 | calcify simp 25 | 26 | /-- 27 | info: Try this: calc 28 | 0 + n 29 | _ = n := Nat.zero_add n 30 | -/ 31 | #guard_msgs in 32 | example (n : Nat) : 0 + n = n := by 33 | calcify simp 34 | 35 | /-- 36 | info: Try this: calc 37 | n 38 | _ = 1 * n := (Nat.one_mul n).symm 39 | -/ 40 | #guard_msgs in 41 | example (n : Nat) : n = 1 * n := by 42 | calcify simp 43 | 44 | /-- 45 | info: Try this: calc 46 | 0 + n 47 | _ = n := (Nat.zero_add n) 48 | _ = 1 * n := (Nat.one_mul n).symm 49 | -/ 50 | #guard_msgs in 51 | example (n : Nat) : 0 + n = 1 * n := by 52 | calcify simp [Nat.zero_add, Nat.one_mul] 53 | 54 | /-- 55 | info: Try this: calc 56 | 0 + n 57 | _ = n := (Nat.zero_add n) 58 | _ = n * 1 := (Nat.mul_one n).symm 59 | _ = 1 * n * 1 := congrArg (fun _a => _a * 1) (Eq.symm (Nat.one_mul n)) 60 | -/ 61 | #guard_msgs in 62 | example (n : Nat) : 0 + n = 1 * n * 1 := by 63 | calcify rw [Nat.zero_add, Nat.one_mul, Nat.mul_one] 64 | 65 | /-- 66 | info: Try this: calc 67 | 0 + n 68 | _ = n := (Nat.zero_add n) 69 | _ = 0 + n := (Nat.zero_add n).symm 70 | _ = 0 + n * 1 := congrArg (fun _a => 0 + _a) (Eq.symm (Nat.mul_one n)) 71 | -/ 72 | #guard_msgs in 73 | example (n : Nat) : 0 + n = 0 + (n * 1) := by 74 | calcify rw [Nat.mul_one, Nat.zero_add] 75 | 76 | /-- 77 | info: Try this: conv => 78 | tactic => 79 | calc 80 | P (0 + 1 * n * 1) 81 | _ = P (0 + n * 1) := (congrArg (fun x => P (0 + x * 1)) (Nat.one_mul n)) 82 | _ = P (0 + n) := (congrArg (fun x => P (0 + x)) (Nat.mul_one n)) 83 | _ = P n := congrArg P (Nat.zero_add n) 84 | -/ 85 | #guard_msgs in 86 | example (n : Nat) (P : Nat → Prop) (h : P n): P (0 + 1 * n * 1) := by 87 | calcify simp 88 | exact h 89 | 90 | /-- 91 | info: Try this: ⏎ 92 | conv => 93 | tactic => 94 | calc 95 | P (0 + 1 * n * 1) 96 | _ = P (0 + n * 1) := (congrArg (fun x => P (0 + x * 1)) (Nat.one_mul n)) 97 | _ = P (0 + n) := (congrArg (fun x => P (0 + x)) (Nat.mul_one n)) 98 | _ = P n := congrArg P (Nat.zero_add n) 99 | refine h 100 | -/ 101 | #guard_msgs in 102 | example (n : Nat) (P : Nat → Prop) (h : P n): P (0 + 1 * n * 1) := by 103 | calcify simp [h] 104 | 105 | 106 | -- Rewriting under binders works: 107 | 108 | /-- 109 | info: Try this: calc 110 | List.map (fun x => (0 + 1) * (0 + x)) xs 111 | _ = List.map (fun x => 1 * (0 + x)) xs := 112 | (congrArg (fun x => List.map x xs) (funext fun n => congrArg (fun x => x * (0 + n)) (Nat.zero_add 1))) 113 | _ = List.map (HMul.hMul 1) xs := 114 | (congrArg (fun x => List.map x xs) (funext fun n => congrArg (HMul.hMul 1) (Nat.zero_add n))) 115 | _ = List.map (fun x => x) xs := (congrArg (fun x => List.map x xs) (funext fun n => Nat.one_mul n)) 116 | _ = id xs := congrArg (fun x => x xs) List.map_id_fun' 117 | -/ 118 | #guard_msgs in 119 | example xs : List.map (fun n => (0 + 1) * (0 + n)) xs = xs := by 120 | calcify simp 121 | 122 | 123 | -- But contextual rewriting using congruence rules are not supported well. 124 | -- We have ad-hoc support for `ite_congr` (see above), but it is far from elegant. 125 | -- Maybe it could be implemented fully generically, but it would be quite hairy I fear. 126 | 127 | /-- 128 | info: Try this: calc 129 | if 0 + 1 * x = 0 then x + (2 * x + n) else 0 + n 130 | _ = if 0 + x = 0 then x + (2 * x + n) else 0 + n := 131 | (ite_congr (congrArg (fun x => 0 + x = 0) (Nat.one_mul x)) (fun a => Eq.refl (x + (2 * x + n))) fun a => 132 | Eq.refl (0 + n)) 133 | _ = if x = 0 then x + (2 * x + n) else 0 + n := 134 | (ite_congr (congrArg (fun x => x = 0) (Nat.zero_add x)) (fun a => Eq.refl (x + (2 * x + n))) fun a => 135 | Eq.refl (0 + n)) 136 | _ = if x = 0 then 0 + (2 * x + n) else 0 + n := 137 | (ite_congr (Eq.refl (x = 0)) (fun a => congrArg (fun x_1 => x_1 + (2 * x + n)) a) fun a => Eq.refl (0 + n)) 138 | _ = if x = 0 then 0 + (2 * 0 + n) else 0 + n := 139 | (ite_congr (Eq.refl (x = 0)) (fun a => congrArg (fun x => 0 + (2 * x + n)) a) fun a => Eq.refl (0 + n)) 140 | _ = if x = 0 then 0 + n else 0 + n := (congrArg (fun x_1 => if x = 0 then 0 + x_1 else 0 + n) (Nat.zero_add n)) 141 | _ = if x = 0 then n else 0 + n := (congrArg (fun x_1 => if x = 0 then x_1 else 0 + n) (Nat.zero_add n)) 142 | _ = if x = 0 then n else n := (congrArg (fun x_1 => if x = 0 then n else x_1) (Nat.zero_add n)) 143 | _ = n := ite_self n 144 | -/ 145 | #guard_msgs in 146 | example (x n : Nat) : (if 0 + (1 * x) = 0 then x + ((2 * x) + n) else 0 + n) = n := by 147 | calcify (simp (config := {contextual := true})) 148 | 149 | /-- 150 | info: Try this: conv => 151 | tactic => 152 | calc 153 | P (if 0 + 1 * x = 0 then x + (2 * x + n) else 0 + n) 154 | _ = P (if 0 + x = 0 then x + (2 * x + n) else 0 + n) := 155 | (congrArg P 156 | (ite_congr (congrArg (fun x => 0 + x = 0) (Nat.one_mul x)) (fun a => Eq.refl (x + (2 * x + n))) fun a => 157 | Eq.refl (0 + n))) 158 | _ = P (if x = 0 then x + (2 * x + n) else 0 + n) := 159 | (congrArg P 160 | (ite_congr (congrArg (fun x => x = 0) (Nat.zero_add x)) (fun a => Eq.refl (x + (2 * x + n))) fun a => 161 | Eq.refl (0 + n))) 162 | _ = P (if x = 0 then 0 + (2 * x + n) else 0 + n) := 163 | (congrArg P 164 | (ite_congr (Eq.refl (x = 0)) (fun a => congrArg (fun x_1 => x_1 + (2 * x + n)) a) fun a => Eq.refl (0 + n))) 165 | _ = P (if x = 0 then 0 + (2 * 0 + n) else 0 + n) := 166 | (congrArg P 167 | (ite_congr (Eq.refl (x = 0)) (fun a => congrArg (fun x => 0 + (2 * x + n)) a) fun a => Eq.refl (0 + n))) 168 | _ = P (if x = 0 then 0 + n else 0 + n) := 169 | (congrArg (fun x_1 => P (if x = 0 then 0 + x_1 else 0 + n)) (Nat.zero_add n)) 170 | _ = P (if x = 0 then n else 0 + n) := (congrArg (fun x_1 => P (if x = 0 then x_1 else 0 + n)) (Nat.zero_add n)) 171 | _ = P (if x = 0 then n else n) := (congrArg (fun x_1 => P (if x = 0 then n else x_1)) (Nat.zero_add n)) 172 | _ = P n := congrArg P (ite_self n) 173 | -/ 174 | #guard_msgs in 175 | example (x n : Nat) (P : Nat → Prop) (hP : P n): P (if 0 + (1 * x) = 0 then x + ((2 * x) + n) else 0 + n) := by 176 | calcify (simp (config := {contextual := true})) 177 | exact hP 178 | 179 | @[congr] 180 | private theorem List.map_congr {f g : α → β} : ∀ {l : List α}, (∀ x ∈ l, f x = g x) → map f l = map g l 181 | | [], _ => rfl 182 | | a :: l, h => by 183 | let ⟨h₁, h₂⟩ := forall_mem_cons.1 h 184 | rw [map, map, h₁, map_congr h₂] 185 | 186 | /-- 187 | info: Try this: calc 188 | List.map (fun n => f n) xs 189 | _ = List.map (fun x => x) xs := 190 | (List.map_congr fun x a => Eq.trans ((fun n a => h n a) x (of_eq_true (eq_true a))) (Nat.one_mul x)) 191 | _ = id xs := congrArg (fun x => x xs) List.map_id_fun' 192 | -/ 193 | #guard_msgs in 194 | example xs (f : Nat → Nat) (h : ∀ n, n ∈ xs → f n = 1 * n) : 195 | List.map (fun n => f n) xs = xs := by 196 | calcify simp (config := {contextual := true}) [h] 197 | -------------------------------------------------------------------------------- /Calcify/Eggsamples.lean: -------------------------------------------------------------------------------- 1 | import Calcify.Basic 2 | 3 | /-- 4 | info: Try this: calc 5 | x 6 | _ = x * x * (x * (x * x)) := (h x (x * x)) 7 | _ = x * x * x := congrArg (fun a'2521 => x * x * a'2521) (Eq.symm (h x x)) 8 | -/ 9 | #guard_msgs in 10 | theorem Equation14_implies_Equation232 (G: Type _) [inst : Mul G] 11 | (h: ∀ x y : G, x = y * (x * y)) : ∀ x : G, x = (x * x) * x := by 12 | intro x 13 | calcify 14 | exact @Eq.trans G x (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x x) (@HMul.hMul G G G (@instHMul G inst) x (@HMul.hMul G G G (@instHMul G inst) x x))) (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x x) x) (h x (@HMul.hMul G G G (@instHMul G inst) x x)) (@eq_of_heq G (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x x) (@HMul.hMul G G G (@instHMul G inst) x (@HMul.hMul G G G (@instHMul G inst) x x))) (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x x) x) ((fun (α : Type u_1) (α' : Type u_1) (e_1 : @Eq (Type u_1) α α') => @Eq.ndrec (Type u_1) α (fun (α' : Type u_1) => forall (β : Type u_1) (β' : Type u_1), (@Eq (Type u_1) β β') -> (forall (γ : Type u_1) (γ' : Type u_1), (@Eq (Type u_1) γ γ') -> (forall (self : HMul.{u_1, u_1, u_1} α β γ) (self' : HMul.{u_1, u_1, u_1} α' β' γ'), (@HEq (HMul.{u_1, u_1, u_1} α β γ) self (HMul.{u_1, u_1, u_1} α' β' γ') self') -> (forall (a2519 : α) (a'2519 : α'), (@HEq α a2519 α' a'2519) -> (forall (a2521 : β) (a'2521 : β'), (@HEq β a2521 β' a'2521) -> (@HEq γ (@HMul.hMul α β γ self a2519 a2521) γ' (@HMul.hMul α' β' γ' self' a'2519 a'2521))))))) (fun (β : Type u_1) (β' : Type u_1) (e_2 : @Eq (Type u_1) β β') => @Eq.ndrec (Type u_1) β (fun (β' : Type u_1) => forall (γ : Type u_1) (γ' : Type u_1), (@Eq (Type u_1) γ γ') -> (forall (self : HMul.{u_1, u_1, u_1} α β γ) (self' : HMul.{u_1, u_1, u_1} α β' γ'), (@HEq (HMul.{u_1, u_1, u_1} α β γ) self (HMul.{u_1, u_1, u_1} α β' γ') self') -> (forall (a2519 : α) (a'2519 : α), (@HEq α a2519 α a'2519) -> (forall (a2521 : β) (a'2521 : β'), (@HEq β a2521 β' a'2521) -> (@HEq γ (@HMul.hMul α β γ self a2519 a2521) γ' (@HMul.hMul α β' γ' self' a'2519 a'2521)))))) (fun (γ : Type u_1) (γ' : Type u_1) (e_3 : @Eq (Type u_1) γ γ') => @Eq.ndrec (Type u_1) γ (fun (γ' : Type u_1) => forall (self : HMul.{u_1, u_1, u_1} α β γ) (self' : HMul.{u_1, u_1, u_1} α β γ'), (@HEq (HMul.{u_1, u_1, u_1} α β γ) self (HMul.{u_1, u_1, u_1} α β γ') self') -> (forall (a2519 : α) (a'2519 : α), (@HEq α a2519 α a'2519) -> (forall (a2521 : β) (a'2521 : β), (@HEq β a2521 β a'2521) -> (@HEq γ (@HMul.hMul α β γ self a2519 a2521) γ' (@HMul.hMul α β γ' self' a'2519 a'2521))))) (fun (self : HMul.{u_1, u_1, u_1} α β γ) (self' : HMul.{u_1, u_1, u_1} α β γ) (e_4 : @HEq (HMul.{u_1, u_1, u_1} α β γ) self (HMul.{u_1, u_1, u_1} α β γ) self') => @Eq.ndrec (HMul.{u_1, u_1, u_1} α β γ) self (fun (self' : HMul.{u_1, u_1, u_1} α β γ) => forall (a2519 : α) (a'2519 : α), (@HEq α a2519 α a'2519) -> (forall (a2521 : β) (a'2521 : β), (@HEq β a2521 β a'2521) -> (@HEq γ (@HMul.hMul α β γ self a2519 a2521) γ (@HMul.hMul α β γ self' a'2519 a'2521)))) (fun (a2519 : α) (a'2519 : α) (e_5 : @HEq α a2519 α a'2519) => @Eq.ndrec α a2519 (fun (a'2519 : α) => forall (a2521 : β) (a'2521 : β), (@HEq β a2521 β a'2521) -> (@HEq γ (@HMul.hMul α β γ self a2519 a2521) γ (@HMul.hMul α β γ self a'2519 a'2521))) (fun (a2521 : β) (a'2521 : β) (e_6 : @HEq β a2521 β a'2521) => @Eq.ndrec β a2521 (fun (a'2521 : β) => @HEq γ (@HMul.hMul α β γ self a2519 a2521) γ (@HMul.hMul α β γ self a2519 a'2521)) (@HEq.refl γ (@HMul.hMul α β γ self a2519 a2521)) a'2521 (@eq_of_heq β a2521 a'2521 e_6)) a'2519 (@eq_of_heq α a2519 a'2519 e_5)) self' (@eq_of_heq (HMul.{u_1, u_1, u_1} α β γ) self self' e_4)) γ' e_3) β' e_2) α' e_1) G G (@Eq.refl (Type u_1) G) G G (@Eq.refl (Type u_1) G) G G (@Eq.refl (Type u_1) G) (@instHMul G inst) (@instHMul G inst) (@HEq.refl (HMul.{u_1, u_1, u_1} G G G) (@instHMul G inst)) (@HMul.hMul G G G (@instHMul G inst) x x) (@HMul.hMul G G G (@instHMul G inst) x x) (@HEq.refl G (@HMul.hMul G G G (@instHMul G inst) x x)) (@HMul.hMul G G G (@instHMul G inst) x (@HMul.hMul G G G (@instHMul G inst) x x)) x (@heq_of_eq G (@HMul.hMul G G G (@instHMul G inst) x (@HMul.hMul G G G (@instHMul G inst) x x)) x (@Eq.symm G x (@HMul.hMul G G G (@instHMul G inst) x (@HMul.hMul G G G (@instHMul G inst) x x)) (h x x))))) 15 | 16 | /-- 17 | info: Try this: calc 18 | x * (y * (y * z * z)) 19 | _ = x * y * x * (x * y * y) * (y * (y * z * z)) := (congrArg (fun a' => a' * (y * (y * z * z))) (h x (x * y) y)) 20 | _ = x * y * x * (x * y * y) * (x * y * y * (y * z * z) * (y * z * z)) := 21 | (congrArg (fun x_1 => x * y * x * (x * y * y) * (x_1 * (y * z * z))) (h y (x * y) z)) 22 | _ = x * y * y := (h (x * y * y) (x * y * x) (y * z * z)).symm 23 | -/ 24 | #guard_msgs in 25 | theorem helper_lemma (G: Type _) [inst : Mul G] (h: ∀ x y z : G, x = (y * x) * ((x * z) * z)) : ∀ x y z : G, x * (y * ((y * z) * z)) = (x * y) * y := by 26 | intro x y z 27 | calcify 28 | exact 29 | @Eq.trans G 30 | (@HMul.hMul G G G (@instHMul G inst) x 31 | (@HMul.hMul G G G (@instHMul G inst) y 32 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z))) 33 | (@HMul.hMul G G G (@instHMul G inst) 34 | (@HMul.hMul G G G (@instHMul G inst) 35 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 36 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y)) 37 | (@HMul.hMul G G G (@instHMul G inst) 38 | (@HMul.hMul G G G (@instHMul G inst) 39 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y) 40 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z)) 41 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z))) 42 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y) 43 | (@Eq.trans G 44 | (@HMul.hMul G G G (@instHMul G inst) x 45 | (@HMul.hMul G G G (@instHMul G inst) y 46 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z))) 47 | (@HMul.hMul G G G (@instHMul G inst) 48 | (@HMul.hMul G G G (@instHMul G inst) 49 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 50 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y)) 51 | (@HMul.hMul G G G (@instHMul G inst) y 52 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z))) 53 | (@HMul.hMul G G G (@instHMul G inst) 54 | (@HMul.hMul G G G (@instHMul G inst) 55 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 56 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y)) 57 | (@HMul.hMul G G G (@instHMul G inst) 58 | (@HMul.hMul G G G (@instHMul G inst) 59 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y) 60 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z)) 61 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z))) 62 | (@eq_of_heq G 63 | (@HMul.hMul G G G (@instHMul G inst) x 64 | (@HMul.hMul G G G (@instHMul G inst) y 65 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z))) 66 | (@HMul.hMul G G G (@instHMul G inst) 67 | (@HMul.hMul G G G (@instHMul G inst) 68 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 69 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y)) 70 | (@HMul.hMul G G G (@instHMul G inst) y 71 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z))) 72 | ((fun α α' e_1 => 73 | @Eq.ndrec (Type u_1) α 74 | (fun α' => 75 | ∀ (β β' : Type u_1), 76 | @Eq (Type u_1) β β' → 77 | ∀ (γ γ' : Type u_1), 78 | @Eq (Type u_1) γ γ' → 79 | ∀ (self : HMul α β γ) (self' : HMul α' β' γ'), 80 | @HEq (HMul α β γ) self (HMul α' β' γ') self' → 81 | ∀ (a : α) (a' : α'), 82 | @HEq α a α' a' → 83 | ∀ (a_1 : β) (a'_1 : β'), 84 | @HEq β a_1 β' a'_1 → 85 | @HEq γ (@HMul.hMul α β γ self a a_1) γ' 86 | (@HMul.hMul α' β' γ' self' a' a'_1)) 87 | (fun β β' e_2 => 88 | @Eq.ndrec (Type u_1) β 89 | (fun β' => 90 | ∀ (γ γ' : Type u_1), 91 | @Eq (Type u_1) γ γ' → 92 | ∀ (self : HMul α β γ) (self' : HMul α β' γ'), 93 | @HEq (HMul α β γ) self (HMul α β' γ') self' → 94 | ∀ (a a' : α), 95 | @HEq α a α a' → 96 | ∀ (a_1 : β) (a'_1 : β'), 97 | @HEq β a_1 β' a'_1 → 98 | @HEq γ (@HMul.hMul α β γ self a a_1) γ' 99 | (@HMul.hMul α β' γ' self' a' a'_1)) 100 | (fun γ γ' e_3 => 101 | @Eq.ndrec (Type u_1) γ 102 | (fun γ' => 103 | ∀ (self : HMul α β γ) (self' : HMul α β γ'), 104 | @HEq (HMul α β γ) self (HMul α β γ') self' → 105 | ∀ (a a' : α), 106 | @HEq α a α a' → 107 | ∀ (a_1 a'_1 : β), 108 | @HEq β a_1 β a'_1 → 109 | @HEq γ (@HMul.hMul α β γ self a a_1) γ' 110 | (@HMul.hMul α β γ' self' a' a'_1)) 111 | (fun self self' e_4 => 112 | @Eq.ndrec (HMul α β γ) self 113 | (fun self' => 114 | ∀ (a a' : α), 115 | @HEq α a α a' → 116 | ∀ (a_1 a'_1 : β), 117 | @HEq β a_1 β a'_1 → 118 | @HEq γ (@HMul.hMul α β γ self a a_1) γ 119 | (@HMul.hMul α β γ self' a' a'_1)) 120 | (fun a a' e_5 => 121 | @Eq.ndrec α a 122 | (fun a' => 123 | ∀ (a_1 a'_1 : β), 124 | @HEq β a_1 β a'_1 → 125 | @HEq γ (@HMul.hMul α β γ self a a_1) γ 126 | (@HMul.hMul α β γ self a' a'_1)) 127 | (fun a_1 a' e_6 => 128 | @Eq.ndrec β a_1 129 | (fun a' => 130 | @HEq γ (@HMul.hMul α β γ self a a_1) γ 131 | (@HMul.hMul α β γ self a a')) 132 | (@HEq.refl γ (@HMul.hMul α β γ self a a_1)) a' 133 | (@eq_of_heq β a_1 a' e_6)) 134 | a' (@eq_of_heq α a a' e_5)) 135 | self' (@eq_of_heq (HMul α β γ) self self' e_4)) 136 | γ' e_3) 137 | β' e_2) 138 | α' e_1) 139 | G G (@Eq.refl (Type u_1) G) G G (@Eq.refl (Type u_1) G) G G (@Eq.refl (Type u_1) G) 140 | (@instHMul G inst) (@instHMul G inst) (@HEq.refl (HMul G G G) (@instHMul G inst)) x 141 | (@HMul.hMul G G G (@instHMul G inst) 142 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 143 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y)) 144 | (@heq_of_eq G x 145 | (@HMul.hMul G G G (@instHMul G inst) 146 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 147 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y)) 148 | (h x (@HMul.hMul G G G (@instHMul G inst) x y) y)) 149 | (@HMul.hMul G G G (@instHMul G inst) y 150 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z)) 151 | (@HMul.hMul G G G (@instHMul G inst) y 152 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z)) 153 | (@HEq.refl G 154 | (@HMul.hMul G G G (@instHMul G inst) y 155 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) 156 | z))))) 157 | (@eq_of_heq G 158 | (@HMul.hMul G G G (@instHMul G inst) 159 | (@HMul.hMul G G G (@instHMul G inst) 160 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 161 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y)) 162 | (@HMul.hMul G G G (@instHMul G inst) y 163 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z))) 164 | (@HMul.hMul G G G (@instHMul G inst) 165 | (@HMul.hMul G G G (@instHMul G inst) 166 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 167 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y)) 168 | (@HMul.hMul G G G (@instHMul G inst) 169 | (@HMul.hMul G G G (@instHMul G inst) 170 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y) 171 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z)) 172 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z))) 173 | ((fun α α' e_1 => 174 | @Eq.ndrec (Type u_1) α 175 | (fun α' => 176 | ∀ (β β' : Type u_1), 177 | @Eq (Type u_1) β β' → 178 | ∀ (γ γ' : Type u_1), 179 | @Eq (Type u_1) γ γ' → 180 | ∀ (self : HMul α β γ) (self' : HMul α' β' γ'), 181 | @HEq (HMul α β γ) self (HMul α' β' γ') self' → 182 | ∀ (a : α) (a' : α'), 183 | @HEq α a α' a' → 184 | ∀ (a_1 : β) (a'_1 : β'), 185 | @HEq β a_1 β' a'_1 → 186 | @HEq γ (@HMul.hMul α β γ self a a_1) γ' 187 | (@HMul.hMul α' β' γ' self' a' a'_1)) 188 | (fun β β' e_2 => 189 | @Eq.ndrec (Type u_1) β 190 | (fun β' => 191 | ∀ (γ γ' : Type u_1), 192 | @Eq (Type u_1) γ γ' → 193 | ∀ (self : HMul α β γ) (self' : HMul α β' γ'), 194 | @HEq (HMul α β γ) self (HMul α β' γ') self' → 195 | ∀ (a a' : α), 196 | @HEq α a α a' → 197 | ∀ (a_1 : β) (a'_1 : β'), 198 | @HEq β a_1 β' a'_1 → 199 | @HEq γ (@HMul.hMul α β γ self a a_1) γ' 200 | (@HMul.hMul α β' γ' self' a' a'_1)) 201 | (fun γ γ' e_3 => 202 | @Eq.ndrec (Type u_1) γ 203 | (fun γ' => 204 | ∀ (self : HMul α β γ) (self' : HMul α β γ'), 205 | @HEq (HMul α β γ) self (HMul α β γ') self' → 206 | ∀ (a a' : α), 207 | @HEq α a α a' → 208 | ∀ (a_1 a'_1 : β), 209 | @HEq β a_1 β a'_1 → 210 | @HEq γ (@HMul.hMul α β γ self a a_1) γ' 211 | (@HMul.hMul α β γ' self' a' a'_1)) 212 | (fun self self' e_4 => 213 | @Eq.ndrec (HMul α β γ) self 214 | (fun self' => 215 | ∀ (a a' : α), 216 | @HEq α a α a' → 217 | ∀ (a_1 a'_1 : β), 218 | @HEq β a_1 β a'_1 → 219 | @HEq γ (@HMul.hMul α β γ self a a_1) γ 220 | (@HMul.hMul α β γ self' a' a'_1)) 221 | (fun a a' e_5 => 222 | @Eq.ndrec α a 223 | (fun a' => 224 | ∀ (a_1 a'_1 : β), 225 | @HEq β a_1 β a'_1 → 226 | @HEq γ (@HMul.hMul α β γ self a a_1) γ 227 | (@HMul.hMul α β γ self a' a'_1)) 228 | (fun a_1 a' e_6 => 229 | @Eq.ndrec β a_1 230 | (fun a' => 231 | @HEq γ (@HMul.hMul α β γ self a a_1) γ 232 | (@HMul.hMul α β γ self a a')) 233 | (@HEq.refl γ (@HMul.hMul α β γ self a a_1)) a' 234 | (@eq_of_heq β a_1 a' e_6)) 235 | a' (@eq_of_heq α a a' e_5)) 236 | self' (@eq_of_heq (HMul α β γ) self self' e_4)) 237 | γ' e_3) 238 | β' e_2) 239 | α' e_1) 240 | G G (@Eq.refl (Type u_1) G) G G (@Eq.refl (Type u_1) G) G G (@Eq.refl (Type u_1) G) 241 | (@instHMul G inst) (@instHMul G inst) (@HEq.refl (HMul G G G) (@instHMul G inst)) 242 | (@HMul.hMul G G G (@instHMul G inst) 243 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 244 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y)) 245 | (@HMul.hMul G G G (@instHMul G inst) 246 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 247 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y)) 248 | (@HEq.refl G 249 | (@HMul.hMul G G G (@instHMul G inst) 250 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 251 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y))) 252 | (@HMul.hMul G G G (@instHMul G inst) y 253 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z)) 254 | (@HMul.hMul G G G (@instHMul G inst) 255 | (@HMul.hMul G G G (@instHMul G inst) 256 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y) 257 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z)) 258 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z)) 259 | ((fun α α' e_1 => 260 | @Eq.ndrec (Type u_1) α 261 | (fun α' => 262 | ∀ (β β' : Type u_1), 263 | @Eq (Type u_1) β β' → 264 | ∀ (γ γ' : Type u_1), 265 | @Eq (Type u_1) γ γ' → 266 | ∀ (self : HMul α β γ) (self' : HMul α' β' γ'), 267 | @HEq (HMul α β γ) self (HMul α' β' γ') self' → 268 | ∀ (a : α) (a' : α'), 269 | @HEq α a α' a' → 270 | ∀ (a_1 : β) (a'_1 : β'), 271 | @HEq β a_1 β' a'_1 → 272 | @HEq γ (@HMul.hMul α β γ self a a_1) γ' 273 | (@HMul.hMul α' β' γ' self' a' a'_1)) 274 | (fun β β' e_2 => 275 | @Eq.ndrec (Type u_1) β 276 | (fun β' => 277 | ∀ (γ γ' : Type u_1), 278 | @Eq (Type u_1) γ γ' → 279 | ∀ (self : HMul α β γ) (self' : HMul α β' γ'), 280 | @HEq (HMul α β γ) self (HMul α β' γ') self' → 281 | ∀ (a a' : α), 282 | @HEq α a α a' → 283 | ∀ (a_1 : β) (a'_1 : β'), 284 | @HEq β a_1 β' a'_1 → 285 | @HEq γ (@HMul.hMul α β γ self a a_1) γ' 286 | (@HMul.hMul α β' γ' self' a' a'_1)) 287 | (fun γ γ' e_3 => 288 | @Eq.ndrec (Type u_1) γ 289 | (fun γ' => 290 | ∀ (self : HMul α β γ) (self' : HMul α β γ'), 291 | @HEq (HMul α β γ) self (HMul α β γ') self' → 292 | ∀ (a a' : α), 293 | @HEq α a α a' → 294 | ∀ (a_1 a'_1 : β), 295 | @HEq β a_1 β a'_1 → 296 | @HEq γ (@HMul.hMul α β γ self a a_1) γ' 297 | (@HMul.hMul α β γ' self' a' a'_1)) 298 | (fun self self' e_4 => 299 | @Eq.ndrec (HMul α β γ) self 300 | (fun self' => 301 | ∀ (a a' : α), 302 | @HEq α a α a' → 303 | ∀ (a_1 a'_1 : β), 304 | @HEq β a_1 β a'_1 → 305 | @HEq γ (@HMul.hMul α β γ self a a_1) γ 306 | (@HMul.hMul α β γ self' a' a'_1)) 307 | (fun a a' e_5 => 308 | @Eq.ndrec α a 309 | (fun a' => 310 | ∀ (a_1 a'_1 : β), 311 | @HEq β a_1 β a'_1 → 312 | @HEq γ (@HMul.hMul α β γ self a a_1) γ 313 | (@HMul.hMul α β γ self a' a'_1)) 314 | (fun a_1 a' e_6 => 315 | @Eq.ndrec β a_1 316 | (fun a' => 317 | @HEq γ (@HMul.hMul α β γ self a a_1) γ 318 | (@HMul.hMul α β γ self a a')) 319 | (@HEq.refl γ (@HMul.hMul α β γ self a a_1)) a' 320 | (@eq_of_heq β a_1 a' e_6)) 321 | a' (@eq_of_heq α a a' e_5)) 322 | self' (@eq_of_heq (HMul α β γ) self self' e_4)) 323 | γ' e_3) 324 | β' e_2) 325 | α' e_1) 326 | G G (@Eq.refl (Type u_1) G) G G (@Eq.refl (Type u_1) G) G G (@Eq.refl (Type u_1) G) 327 | (@instHMul G inst) (@instHMul G inst) (@HEq.refl (HMul G G G) (@instHMul G inst)) y 328 | (@HMul.hMul G G G (@instHMul G inst) 329 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y) 330 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z)) 331 | (@heq_of_eq G y 332 | (@HMul.hMul G G G (@instHMul G inst) 333 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y) 334 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z)) 335 | (h y (@HMul.hMul G G G (@instHMul G inst) x y) z)) 336 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z) 337 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z) 338 | (@HEq.refl G 339 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) 340 | z)))))) 341 | (@Eq.symm G (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y) 342 | (@HMul.hMul G G G (@instHMul G inst) 343 | (@HMul.hMul G G G (@instHMul G inst) 344 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 345 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y)) 346 | (@HMul.hMul G G G (@instHMul G inst) 347 | (@HMul.hMul G G G (@instHMul G inst) 348 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y) 349 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z)) 350 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z))) 351 | (h (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) y) 352 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) x y) x) 353 | (@HMul.hMul G G G (@instHMul G inst) (@HMul.hMul G G G (@instHMul G inst) y z) z))) 354 | -------------------------------------------------------------------------------- /Calcify/Tests.lean: -------------------------------------------------------------------------------- 1 | import Calcify 2 | import Calcify.Demo 3 | 4 | section issue9 5 | 6 | axiom coe : Nat → Int 7 | axiom cast_mul : (m n : Nat) → coe (m * n) = coe m * coe n 8 | axiom Gamma : Int → Int 9 | axiom fact: Nat → Nat 10 | axiom Gamma_nat_eq_factorial (n : Nat) : Gamma (coe n + 1) = coe (fact n) 11 | 12 | 13 | /-- 14 | info: Try this: calc 15 | Gamma (coe n + 1) * Gamma (coe n + 1) 16 | _ = coe (fact n) * Gamma (coe n + 1) := (congrArg (fun x => x * Gamma (coe n + 1)) (Gamma_nat_eq_factorial n)) 17 | _ = coe (fact n) * coe (fact n) := (congrArg (fun x => coe (fact n) * x) (Gamma_nat_eq_factorial n)) 18 | _ = coe (fact n * fact n) := (cast_mul (fact n) (fact n)).symm 19 | _ = coe (fact n) * coe (fact n) := cast_mul (fact n) (fact n) 20 | -/ 21 | #guard_msgs in 22 | example (n : Nat) : Gamma (coe n + 1) * Gamma (coe n + 1) = coe (fact n) * coe (fact n) := by 23 | calcify simp only [← cast_mul, Gamma_nat_eq_factorial] 24 | 25 | end issue9 26 | 27 | section issue10 28 | 29 | /-- 30 | info: Try this: ⏎ 31 | apply of_eq_true 32 | calc 33 | True ∧ True 34 | _ = True := and_self True 35 | -/ 36 | #guard_msgs in 37 | example : True ∧ True := by 38 | calcify simp 39 | 40 | 41 | /-- 42 | info: Try this: calc 43 | True ∧ True ∧ True 44 | _ = (True ∧ True) := (congrArg (fun x => True ∧ x) (and_self True)) 45 | _ = True := (and_self True) 46 | _ ↔ True := Iff.rfl 47 | -/ 48 | #guard_msgs in 49 | example : (True ∧ True ∧ True) ↔ True := by 50 | calcify simp 51 | 52 | /-- 53 | info: Try this: calc 54 | True ∧ False ∧ True 55 | _ = (True ∧ False) := (congrArg (fun x => True ∧ x) (and_true False)) 56 | _ = False := (and_false True) 57 | _ = (False ∧ True) := (and_true False).symm 58 | _ ↔ False ∧ True := Iff.rfl 59 | -/ 60 | #guard_msgs in 61 | example : (True ∧ False ∧ True) ↔ (False ∧ True) := by 62 | calcify simp 63 | 64 | /-- 65 | info: Try this: calc 66 | True ∧ True ∧ True 67 | _ = (True ∧ True) := (congrArg (fun x => True ∧ x) (and_self True)) 68 | _ = True := and_self True 69 | -/ 70 | #guard_msgs in 71 | example : (True ∧ True ∧ True) = True := by 72 | calcify simp 73 | 74 | /-- info: Try this: exact Iff.refl True -/ 75 | #guard_msgs in 76 | example : True ↔ True := by 77 | calcify simp 78 | 79 | /-- info: Try this: exact True.intro -/ 80 | #guard_msgs in 81 | example : True := by 82 | calcify simp 83 | 84 | 85 | end issue10 86 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Apache License 2.0 (Apache) 2 | Apache License 3 | Version 2.0, January 2004 4 | http://www.apache.org/licenses/ 5 | 6 | TERMS AND CONDITIONS FOR USE, REPRODUCTION, AND DISTRIBUTION 7 | 8 | 1. Definitions. 9 | 10 | "License" shall mean the terms and conditions for use, reproduction, and distribution as defined by Sections 1 through 9 of this document. 11 | 12 | "Licensor" shall mean the copyright owner or entity authorized by the copyright owner that is granting the License. 13 | 14 | "Legal Entity" shall mean the union of the acting entity and all other entities that control, are controlled by, or are under common control with that entity. For the purposes of this definition, "control" means (i) the power, direct or indirect, to cause the direction or management of such entity, whether by contract or otherwise, or (ii) ownership of fifty percent (50%) or more of the outstanding shares, or (iii) beneficial ownership of such entity. 15 | 16 | "You" (or "Your") shall mean an individual or Legal Entity exercising permissions granted by this License. 17 | 18 | "Source" form shall mean the preferred form for making modifications, including but not limited to software source code, documentation source, and configuration files. 19 | 20 | "Object" form shall mean any form resulting from mechanical transformation or translation of a Source form, including but not limited to compiled object code, generated documentation, and conversions to other media types. 21 | 22 | "Work" shall mean the work of authorship, whether in Source or Object form, made available under the License, as indicated by a copyright notice that is included in or attached to the work (an example is provided in the Appendix below). 23 | 24 | "Derivative Works" shall mean any work, whether in Source or Object form, that is based on (or derived from) the Work and for which the editorial revisions, annotations, elaborations, or other modifications represent, as a whole, an original work of authorship. For the purposes of this License, Derivative Works shall not include works that remain separable from, or merely link (or bind by name) to the interfaces of, the Work and Derivative Works thereof. 25 | 26 | "Contribution" shall mean any work of authorship, including the original version of the Work and any modifications or additions to that Work or Derivative Works thereof, that is intentionally submitted to Licensor for inclusion in the Work by the copyright owner or by an individual or Legal Entity authorized to submit on behalf of the copyright owner. For the purposes of this definition, "submitted" means any form of electronic, verbal, or written communication sent to the Licensor or its representatives, including but not limited to communication on electronic mailing lists, source code control systems, and issue tracking systems that are managed by, or on behalf of, the Licensor for the purpose of discussing and improving the Work, but excluding communication that is conspicuously marked or otherwise designated in writing by the copyright owner as "Not a Contribution." 27 | 28 | "Contributor" shall mean Licensor and any individual or Legal Entity on behalf of whom a Contribution has been received by Licensor and subsequently incorporated within the Work. 29 | 30 | 2. Grant of Copyright License. 31 | 32 | Subject to the terms and conditions of this License, each Contributor hereby grants to You a perpetual, worldwide, non-exclusive, no-charge, royalty-free, irrevocable copyright license to reproduce, prepare Derivative Works of, publicly display, publicly perform, sublicense, and distribute the Work and such Derivative Works in Source or Object form. 33 | 34 | 3. Grant of Patent License. 35 | 36 | Subject to the terms and conditions of this License, each Contributor hereby grants to You a perpetual, worldwide, non-exclusive, no-charge, royalty-free, irrevocable (except as stated in this section) patent license to make, have made, use, offer to sell, sell, import, and otherwise transfer the Work, where such license applies only to those patent claims licensable by such Contributor that are necessarily infringed by their Contribution(s) alone or by combination of their Contribution(s) with the Work to which such Contribution(s) was submitted. If You institute patent litigation against any entity (including a cross-claim or counterclaim in a lawsuit) alleging that the Work or a Contribution incorporated within the Work constitutes direct or contributory patent infringement, then any patent licenses granted to You under this License for that Work shall terminate as of the date such litigation is filed. 37 | 38 | 4. Redistribution. 39 | 40 | You may reproduce and distribute copies of the Work or Derivative Works thereof in any medium, with or without modifications, and in Source or Object form, provided that You meet the following conditions: 41 | 42 | 1. You must give any other recipients of the Work or Derivative Works a copy of this License; and 43 | 44 | 2. You must cause any modified files to carry prominent notices stating that You changed the files; and 45 | 46 | 3. You must retain, in the Source form of any Derivative Works that You distribute, all copyright, patent, trademark, and attribution notices from the Source form of the Work, excluding those notices that do not pertain to any part of the Derivative Works; and 47 | 48 | 4. If the Work includes a "NOTICE" text file as part of its distribution, then any Derivative Works that You distribute must include a readable copy of the attribution notices contained within such NOTICE file, excluding those notices that do not pertain to any part of the Derivative Works, in at least one of the following places: within a NOTICE text file distributed as part of the Derivative Works; within the Source form or documentation, if provided along with the Derivative Works; or, within a display generated by the Derivative Works, if and wherever such third-party notices normally appear. The contents of the NOTICE file are for informational purposes only and do not modify the License. You may add Your own attribution notices within Derivative Works that You distribute, alongside or as an addendum to the NOTICE text from the Work, provided that such additional attribution notices cannot be construed as modifying the License. 49 | 50 | You may add Your own copyright statement to Your modifications and may provide additional or different license terms and conditions for use, reproduction, or distribution of Your modifications, or for any such Derivative Works as a whole, provided Your use, reproduction, and distribution of the Work otherwise complies with the conditions stated in this License. 51 | 52 | 5. Submission of Contributions. 53 | 54 | Unless You explicitly state otherwise, any Contribution intentionally submitted for inclusion in the Work by You to the Licensor shall be under the terms and conditions of this License, without any additional terms or conditions. Notwithstanding the above, nothing herein shall supersede or modify the terms of any separate license agreement you may have executed with Licensor regarding such Contributions. 55 | 56 | 6. Trademarks. 57 | 58 | This License does not grant permission to use the trade names, trademarks, service marks, or product names of the Licensor, except as required for reasonable and customary use in describing the origin of the Work and reproducing the content of the NOTICE file. 59 | 60 | 7. Disclaimer of Warranty. 61 | 62 | Unless required by applicable law or agreed to in writing, Licensor provides the Work (and each Contributor provides its Contributions) on an "AS IS" BASIS, WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied, including, without limitation, any warranties or conditions of TITLE, NON-INFRINGEMENT, MERCHANTABILITY, or FITNESS FOR A PARTICULAR PURPOSE. You are solely responsible for determining the appropriateness of using or redistributing the Work and assume any risks associated with Your exercise of permissions under this License. 63 | 64 | 8. Limitation of Liability. 65 | 66 | In no event and under no legal theory, whether in tort (including negligence), contract, or otherwise, unless required by applicable law (such as deliberate and grossly negligent acts) or agreed to in writing, shall any Contributor be liable to You for damages, including any direct, indirect, special, incidental, or consequential damages of any character arising as a result of this License or out of the use or inability to use the Work (including but not limited to damages for loss of goodwill, work stoppage, computer failure or malfunction, or any and all other commercial damages or losses), even if such Contributor has been advised of the possibility of such damages. 67 | 68 | 9. Accepting Warranty or Additional Liability. 69 | 70 | While redistributing the Work or Derivative Works thereof, You may choose to offer, and charge a fee for, acceptance of support, warranty, indemnity, or other liability obligations and/or rights consistent with this License. However, in accepting such obligations, You may act only on Your own behalf and on Your sole responsibility, not on behalf of any other Contributor, and only if You agree to indemnify, defend, and hold each Contributor harmless for any liability incurred by, or claims asserted against, such Contributor by reason of your accepting any such warranty or additional liability. 71 | 72 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | calcify - an explanation tactic for Lean 2 | ======================================== 3 | 4 | The `calcify` tactic in this repository will take a proof as generated by 5 | tactics like `simp` and `rw`, and shows an equivalent proof by small-step 6 | equational reasoning using `calc`: 7 | 8 | ``` 9 | example (n : Nat) : 0 + n = 0 + (n * 1) := by 10 | calcify rw [Nat.mul_one, Nat.zero_add] 11 | /- 12 | Try this: calc 13 | 0 + n 14 | _ = n := (Nat.zero_add n) 15 | _ = 0 + n := (Eq.symm (Nat.zero_add n)) 16 | _ = 0 + n * 1 := congrArg (fun _a => 0 + _a) (Eq.symm (Nat.mul_one n)) 17 | -/ 18 | ``` 19 | 20 | For more examples, see [`Calcify/Basic.lean`](Calcify/Basic.lean). 21 | 22 | It is mostly an experiment and of prototype quality, but maybe you'll enjoy it. 23 | 24 | Usage 25 | ----- 26 | 27 | This isn’t well tested, but if you are on a new enough version of lean (a 28 | nightly release most likely, or 0.4.7 once it is there), you should be able to 29 | add this repository as a lake dependency and then `import Calcify`. 30 | 31 | 32 | Implementation 33 | -------------- 34 | 35 | The tactic first rewrites the proof term, pushing congruences into `Eq.trans`, 36 | so that the proof term becomes a list of `Eq.trans` applications. It is using a 37 | hand-rolled simplification routine because Lean’s `simp` does not work on proof 38 | terms (as usually they do not matter due to proof irrelvance). 39 | 40 | Then that proof term is turned into a `calc` proof. 41 | 42 | Contact 43 | ------- 44 | 45 | This was created by Joachim Breitner. Feel free to report issues, create PRs 46 | with test cases and more features, or [reach out on 47 | Zulip](https://leanprover.zulipchat.com/#user/470149). 48 | 49 | 50 | 51 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": [], 4 | "name": "calcify", 5 | "lakeDir": ".lake"} 6 | -------------------------------------------------------------------------------- /lakefile.toml: -------------------------------------------------------------------------------- 1 | name = "calcify" 2 | defaultTargets = ["Calcify"] 3 | 4 | [[lean_lib]] 5 | name = "Calcify" 6 | 7 | [[lean_lib]] 8 | name = "tests" 9 | globs = "Calcify.Tests" 10 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.16.0 2 | --------------------------------------------------------------------------------