├── .gitignore ├── lean-toolchain ├── lake-manifest.json ├── lakefile.lean ├── LICENSE ├── MutualInductionTest ├── Tree.lean ├── TypeTheory.lean ├── EvenOdd.lean └── Ordinals.lean ├── AllTest └── All.lean ├── JointTest ├── TypeTheory.lean ├── Ordinals.lean └── EvenOdd.lean ├── Joint.lean ├── MkAll.lean ├── MutualInduction.lean └── README.md /.gitignore: -------------------------------------------------------------------------------- 1 | .lake/ 2 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.25.0-rc2 -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": [], 4 | "name": "MutualInduction", 5 | "lakeDir": ".lake"} 6 | -------------------------------------------------------------------------------- /lakefile.lean: -------------------------------------------------------------------------------- 1 | import Lake 2 | 3 | open Lake DSL 4 | 5 | package «MutualInduction» where 6 | 7 | version := v!"0.1.0" 8 | description := "A mutual induction tactic for Lean 4." 9 | license := "Zlib" 10 | reservoir := true 11 | 12 | @[default_target] 13 | lean_lib «MutualInduction» where 14 | roots := #[`MutualInduction, `Joint] 15 | leanOptions := #[ 16 | ⟨`autoImplicit, false⟩, 17 | ⟨`experimental.module, true⟩ 18 | ] 19 | 20 | @[default_target] 21 | lean_lib «MkAll» where 22 | leanOptions := #[ 23 | ⟨`autoImplicit, false⟩, 24 | ⟨`experimental.module, true⟩ 25 | ] 26 | 27 | @[test_driver] 28 | lean_lib «Test» where 29 | globs := #[`MutualInductionTest.+, `JointTest.+, `AllTest.+] 30 | leanOptions := #[ 31 | ⟨`autoImplicit, false⟩, 32 | ⟨`pp.fieldNotation, false⟩, 33 | ⟨`pp.proofs, true⟩, 34 | ⟨`experimental.module, true⟩ 35 | ] 36 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | zlib License 2 | 3 | Copyright (c) 2025 Jonathan Chan 4 | 5 | This software is provided 'as-is', without any express or implied warranty. 6 | In no event will the authors be held liable for any damages arising from the use of this software. 7 | 8 | Permission is granted to anyone to use this software for any purpose, 9 | including commercial applications, and to alter it and redistribute it freely, 10 | subject to the following restrictions: 11 | 12 | 1. The origin of this software must not be misrepresented; 13 | you must not claim that you wrote the original software. 14 | If you use this software in a product, 15 | an acknowledgment in the product documentation 16 | would be appreciated but is not required. 17 | 2. Altered source versions must be plainly marked as such, 18 | and must not be misrepresented as being the original software. 19 | 3. This notice may not be removed or altered from any source distribution. 20 | -------------------------------------------------------------------------------- /MutualInductionTest/Tree.lean: -------------------------------------------------------------------------------- 1 | import MutualInduction 2 | 3 | mutual 4 | inductive Tree : Prop where 5 | | node : Forest → Tree 6 | inductive Forest : Prop where 7 | | nil : Forest 8 | | cons : Tree → Forest → Forest 9 | end 10 | open Tree Forest 11 | 12 | mutual 13 | inductive AllTree (P : Tree → Prop) (Q : Forest → Prop) : Tree → Prop where 14 | | node : ∀ {f}, Q f → AllForest P Q f → AllTree P Q (node f) 15 | inductive AllForest (P : Tree → Prop) (Q : Forest → Prop) : Forest → Prop where 16 | | nil : AllForest P Q nil 17 | | cons : ∀ {t f}, P t → Q f → AllTree P Q t → AllForest P Q f → AllForest P Q (cons t f) 18 | end 19 | 20 | theorem elim (P : Tree → Prop) (Q : Forest → Prop) 21 | (ht : ∀ {t}, AllTree P Q t → P t) 22 | (hf : ∀ {f}, AllForest P Q f → Q f) : 23 | (∀ t, P t) ∧ (∀ f, Q f) := by 24 | constructor 25 | case' left => intro t; apply ht 26 | case' right => intro f; apply hf 27 | mutual_induction f, t 28 | case node f ihf => exact AllTree.node (hf ihf) ihf 29 | case nil => exact AllForest.nil 30 | case cons t f iht ihf => exact AllForest.cons (ht iht) (hf ihf) iht ihf 31 | 32 | inductive RoseTree (α : Type) : Type where 33 | | node : α → List (RoseTree α) → RoseTree α 34 | open RoseTree 35 | 36 | inductive List.All {α} (P : α → Prop) : List α → Prop where 37 | | nil : All P [] 38 | | cons : ∀ {x xs}, P x → All P xs → All P (x :: xs) 39 | 40 | theorem RoseTree.elim {α} (P : RoseTree α → Prop) 41 | (hnode : ∀ {x t}, List.All P t → P (node x t)) : ∀ t, P t := by 42 | intro t 43 | induction t using RoseTree.rec (motive_2 := List.All P) 44 | case node ih => exact hnode ih 45 | case nil => constructor 46 | case cons => constructor <;> assumption 47 | 48 | inductive RoseTree.belowAll {α} {motive : RoseTree α → Prop} : RoseTree α → Prop where 49 | | node : ∀ {x t}, List.All motive t → List.All belowAll t → belowAll (node x t) 50 | 51 | theorem RoseTree.brecOnAll {α} {P : RoseTree α → Prop} 52 | (hnode : ∀ {t}, RoseTree.belowAll (motive := P) t → P t) : 53 | ∀ t, P t := by 54 | intro t 55 | apply hnode 56 | induction t using RoseTree.elim 57 | case _ ts => 58 | constructor 59 | induction ts <;> constructor 60 | apply hnode; assumption; assumption; assumption 61 | -------------------------------------------------------------------------------- /AllTest/All.lean: -------------------------------------------------------------------------------- 1 | import MkAll 2 | import MutualInduction 3 | 4 | @[mk_all α β] 5 | inductive Fan (α β : Type) : Type where 6 | | nil : Fan α β 7 | | cons : (Nat → α) → (Nat → β) → (Nat → Fan α β) → Fan α β 8 | open Fan 9 | 10 | inductive Tree (α : Type) : Type where 11 | | node : α → Fan (Tree α) (Tree α) → Tree α 12 | open Tree 13 | 14 | @[simp] 15 | noncomputable def Tree.elim {α} {P : Tree α → Type} 16 | (node : ∀ {x t}, Fan.all P P t → P (node x t)) : ∀ t, P t := by 17 | intro t 18 | induction t using Tree.rec (motive_2 := Fan.all P P) 19 | all_goals first | (repeat' constructor) <;> assumption | apply_rules 20 | -- Tree.rec (motive_2 := Fan.all P) 21 | -- (λ _ _ ↦ hnode) ⟨⟩ (λ _ _ ihα iht ↦ ⟨ihα, iht⟩) 22 | 23 | theorem Tree.ielim {α} {P : Tree α → Prop} 24 | (node : ∀ {x t}, Fan.iall P P t → P (node x t)) : ∀ t, P t := by 25 | intro t 26 | induction t using Tree.rec (motive_2 := Fan.iall P P) 27 | all_goals first | (repeat' constructor) <;> assumption | apply_rules 28 | 29 | @[mk_all α] 30 | inductive IFan (α : Prop) : Prop where 31 | | nil : IFan α 32 | | cons : (Nat → α) → (Nat → IFan α) → IFan α 33 | open IFan 34 | 35 | inductive ITree (α : Type) : Prop where 36 | | node : α → IFan (ITree α) → ITree α 37 | open ITree 38 | 39 | @[induction_eliminator] 40 | theorem ITree.elim {α} {P : ITree α → Prop} 41 | (node : ∀ {x t}, IFan.all P t → P (node x t)) : ∀ t, P t := by 42 | intro t 43 | induction t using ITree.rec (motive_2 := IFan.all P) 44 | all_goals first | (repeat' constructor) <;> assumption | apply_rules 45 | 46 | @[mk_all α] 47 | inductive Vec (α : Type) : Nat → Type where 48 | | nil : Vec α 0 49 | | cons : ∀ n, α → Vec α n → Vec α (n + 1) 50 | 51 | inductive VTree (α : Type) : Type where 52 | | node : ∀ n, α → Vec (VTree α) n → VTree α 53 | 54 | theorem VTree.ielim {α} {P : VTree α → Prop} 55 | (node : ∀ {n x t}, Vec.iall P t → P (node n x t)) : ∀ t, P t := by 56 | intro t 57 | induction t using VTree.rec (motive_2 := λ _ ↦ Vec.iall P) 58 | all_goals first | (repeat' constructor) <;> assumption | apply_rules 59 | 60 | @[induction_eliminator] 61 | noncomputable def VTree.elim {α} {P : VTree α → Type} 62 | (node : ∀ {n x t}, Vec.all P t → P (node n x t)) : ∀ t, P t := by 63 | intro t 64 | induction t using VTree.rec (motive_2 := λ _ ↦ Vec.all P) 65 | all_goals first | (repeat' constructor) <;> assumption | apply_rules 66 | 67 | noncomputable def VTree.sum (vt : VTree Nat) : Nat := by 68 | induction vt 69 | case node k vec all => 70 | induction vec 71 | case nil => exact k 72 | case cons vt ih => 73 | simp [Vec.all] at all 74 | let ⟨n, m⟩ := all 75 | exact n + (ih m) 76 | -------------------------------------------------------------------------------- /JointTest/TypeTheory.lean: -------------------------------------------------------------------------------- 1 | import MutualInduction 2 | import Joint 3 | 4 | inductive Term : Type where 5 | | var : Nat → Term 6 | | type : Term 7 | | arr : Term → Term → Term 8 | | abs : Term → Term 9 | | app : Term → Term → Term 10 | 11 | inductive Ctxt : Type where 12 | | nil : Ctxt 13 | | cons : Ctxt → Term → Ctxt 14 | notation:50 "⬝" => Ctxt.nil 15 | infixl:50 "∷" => Ctxt.cons 16 | 17 | axiom subst : Term → Term → Term 18 | axiom In : Nat → Term → Ctxt → Prop 19 | axiom Eqv : Term → Term → Prop 20 | axiom refl : ∀ {a}, Eqv a a 21 | axiom trans : ∀ {a b c}, Eqv a b → Eqv b c → Eqv a c 22 | notation:40 Γ:41 "∋" x:41 "∶" A:41 => In x A Γ 23 | notation:40 a:41 "≡" b:41 => Eqv a b 24 | 25 | section 26 | set_option hygiene false 27 | local notation:40 "⊢" Γ:40 => Wf Γ 28 | local notation:40 Γ:41 "⊢" a:41 "∶" A:41 => Wt Γ a A 29 | open Term 30 | 31 | mutual 32 | -- ⊢ Γ 33 | inductive Wf : Ctxt → Prop where 34 | | nil : ⊢ ⬝ 35 | | cons {Γ A} : 36 | ⊢ Γ → 37 | Γ ⊢ A ∶ type → 38 | --------------- 39 | ⊢ Γ ∷ A 40 | 41 | -- Γ ⊢ a ∶ A 42 | inductive Wt : Ctxt → Term → Term → Prop where 43 | | var {Γ x A} : 44 | Γ ∋ x ∶ A → 45 | ⊢ Γ → 46 | --------------- 47 | Γ ⊢ var x ∶ A 48 | | type {Γ} : 49 | ⊢ Γ → 50 | ----------------- 51 | Γ ⊢ type ∶ type 52 | | arr {Γ A B} : 53 | Γ ⊢ A ∶ type → 54 | Γ ∷ A ⊢ B ∶ type → 55 | -------------------- 56 | Γ ⊢ arr A B ∶ type 57 | | abs {Γ b A B} : 58 | Γ ⊢ A ∶ type → 59 | Γ ∷ A ⊢ b ∶ B → 60 | --------------------- 61 | Γ ⊢ abs b ∶ arr A B 62 | | app {Γ b a A B} : 63 | Γ ⊢ b ∶ arr A B → 64 | Γ ⊢ a ∶ A → 65 | ------------------------ 66 | Γ ⊢ app b a ∶ subst B a 67 | | conv {Γ a A B} : 68 | Γ ⊢ a ∶ A → 69 | Γ ⊢ B ∶ type → 70 | A ≡ B → 71 | --------------- 72 | Γ ⊢ a ∶ B 73 | end 74 | end 75 | 76 | notation:40 "⊢" Γ:40 => Wf Γ 77 | notation:40 Γ:41 "⊢" a:41 "∶" A:41 => Wt Γ a A 78 | open Wf Wt 79 | 80 | /- 81 | Recursors generated for ⊢ Γ and Γ ⊢ a ∶ A: 82 | 83 | Wf.rec : 84 | ∀ {motive_1 : ∀ Γ, ⊢ Γ → Prop} 85 | {motive_2 : ∀ Γ a A, Γ ⊢ a ∶ A → Prop}, 86 | ... Wf cases ... 87 | (∀ {Γ x a A} (mem : Γ ∋ x ∶ A) (h : ⊢ Γ), 88 | motive_1 Γ h → 89 | motive_2 Γ (.var x) A (Wt.var mem h)) → 90 | ... other Wt cases ... 91 | ∀ {Γ} (h : ⊢ Γ), motive_1 Γ h 92 | 93 | Wt.rec : 94 | ∀ {motive_1 : ∀ Γ, ⊢ Γ → Prop} 95 | {motive_2 : ∀ Γ a A, Γ ⊢ a ∶ A → Prop}, 96 | ... cases ... 97 | ∀ {Γ a A} (h : Γ ⊢ a ∶ A), motive_2 Γ a A h 98 | -/ 99 | 100 | joint 101 | -- motives 102 | (Q : ∀ {Γ}, ⊢ Γ → Prop) 103 | (P : ∀ {Γ a A}, Γ ⊢ a ∶ A → Prop) 104 | -- Wf subgoals 105 | (nil : Q nil) 106 | (cons : ∀ {Γ A} 107 | (h : ⊢ Γ) 108 | (hA : Γ ⊢ A ∶ .type), 109 | Q h → P hA → Q (cons h hA)) 110 | -- Wt subgoals 111 | (var : ∀ {Γ x A} 112 | (mem : Γ ∋ x ∶ A) 113 | (h : ⊢ Γ), 114 | Q h → P (var mem h)) 115 | (type : ∀ {Γ} 116 | (h : ⊢ Γ), 117 | Q h → P (type h)) 118 | (arr : ∀ {Γ A B} 119 | (hA : Γ ⊢ A ∶ .type) 120 | (hB : Γ ∷ A ⊢ B ∶ .type), 121 | P hA → P hB → P (arr hA hB)) 122 | (abs : ∀ {Γ b A B} 123 | (hA : Γ ⊢ A ∶ .type) 124 | (hb : Γ ∷ A ⊢ b ∶ B), 125 | P hA → P hb → P (abs hA hb)) 126 | (app : ∀ {Γ b a A B} 127 | (hb : Γ ⊢ b ∶ .arr A B) 128 | (ha : Γ ⊢ a ∶ A), 129 | P hb → P ha → P (app hb ha)) 130 | (conv : ∀ {Γ a A B} 131 | (ha : Γ ⊢ a ∶ A) 132 | (hB : Γ ⊢ B ∶ .type) 133 | (e : A ≡ B), 134 | P ha → P hB → P (conv ha hB e)) 135 | -- goal 136 | theorem wfInd {Γ} (h : ⊢ Γ) : Q h 137 | theorem wtInd {Γ a A} (h : Γ ⊢ a ∶ A) : P h 138 | by 139 | mutual_induction h, h 140 | case nil => exact nil 141 | case cons => apply cons <;> assumption 142 | case var => apply var <;> assumption 143 | case type => apply type <;> assumption 144 | case arr => apply arr <;> assumption 145 | case abs => apply abs <;> assumption 146 | case app => apply app <;> assumption 147 | case conv => apply conv <;> assumption 148 | 149 | theorem wtwf {Γ a A} (h : Γ ⊢ a ∶ A) : ⊢ Γ := by 150 | have _ (h : ⊢ Γ) : True := ?dummy 151 | mutual_induction h, h 152 | all_goals try simp at * <;> assumption 153 | -- induction h using Wt.rec (motive_1 := λ _ _ ↦ True) 154 | 155 | joint {Γ : Ctxt} 156 | theorem regWt x A (mem : Γ ∋ x ∶ A) (h : ⊢ Γ) : Γ ⊢ A ∶ .type 157 | theorem regularity a A (h : Γ ⊢ a ∶ A) : Γ ⊢ A ∶ .type 158 | by 159 | mutual_induction h, h 160 | case nil => sorry 161 | case cons => sorry 162 | case var mem _ ih => exact ih _ _ mem 163 | case type => constructor; assumption 164 | case arr => constructor; apply wtwf; assumption 165 | case abs => constructor <;> assumption 166 | case app => sorry 167 | case conv => assumption 168 | -------------------------------------------------------------------------------- /JointTest/Ordinals.lean: -------------------------------------------------------------------------------- 1 | import MutualInduction 2 | import Joint 3 | 4 | mutual 5 | inductive Ordinal : Type where 6 | | succ : Ordinal → Ordinal 7 | | limit : Ordinals → Ordinal 8 | 9 | inductive Ordinals : Type where 10 | | nil : Ordinals 11 | | cons : Ordinal → Ordinals → Ordinals 12 | end 13 | open Ordinal Ordinals 14 | 15 | mutual 16 | inductive Le : Ordinal → Ordinal → Prop where 17 | | succ {i j} : Le i j → Le (.succ i) (.succ j) 18 | | under {i js} : LeUnder i js → Le i (.limit js) 19 | | over {is j} : LeOver is j → Le (.limit is) j 20 | 21 | inductive LeUnder : Ordinal → Ordinals → Prop where 22 | | underHere {i j js} : Le i j → LeUnder i (.cons j js) 23 | | underThere {i j js} : LeUnder i js → LeUnder i (.cons j js) 24 | 25 | inductive LeOver : Ordinals → Ordinal → Prop where 26 | | overNil {j} : LeOver .nil j 27 | | overCons {i is j} : Le i j → LeOver is j → LeOver (.cons i is) j 28 | end 29 | open LeUnder LeOver 30 | 31 | def zero := Ordinal.limit .nil 32 | theorem zeroLe {i} : Le zero i := Le.over .overNil 33 | theorem noLeZero {i} : Le (.succ i) zero → False := by 34 | intro le; cases le 35 | case under leUnder => cases leUnder 36 | 37 | joint 38 | -- motives 39 | (P : ∀ {i j}, Le i j → Prop) 40 | (Q : ∀ {i js}, LeUnder i js → Prop) 41 | (R : ∀ {is j}, LeOver is j → Prop) 42 | -- Le subgoals 43 | (succ : ∀ {i j} 44 | (h : Le i j), 45 | P h → P (.succ h)) 46 | (under : ∀ {i js} 47 | (h : LeUnder i js), 48 | Q h → P (.under h)) 49 | (over : ∀ {is j} 50 | (h : LeOver is j), 51 | R h → P (.over h)) 52 | -- LeUnder subgoals 53 | (underHere : ∀ {i j js} 54 | (h : Le i j), 55 | P h → Q (.underHere (js := js) h)) 56 | (underThere : ∀ {i j js} 57 | (h : LeUnder i js), 58 | Q h → Q (.underThere (j := j) h)) 59 | -- LeOver subgoals 60 | (overNil : ∀ {j}, R (.overNil (j := j))) 61 | (overCons : ∀ {i is j} 62 | (hle : Le i j) 63 | (h : LeOver is j), 64 | P hle → R h → R (.overCons hle h)) 65 | theorem Le.ind {i j} (le : Le i j) : P le 66 | theorem LeUnder.ind {i js} (leunder : LeUnder i js) : Q leunder 67 | theorem LeOver.ind {is j} (leover : LeOver is j) : R leover 68 | by 69 | mutual_induction le, leunder, leover 70 | case underThere ih => apply underThere _ ih 71 | all_goals apply_rules 72 | 73 | @[simp] 74 | def succOrdinals : Ordinals → Ordinals 75 | | nil => nil 76 | | cons i is => cons (succ i) (succOrdinals is) 77 | 78 | joint 79 | theorem leSuccInv {i j} (h : Le (succ i) (succ j)) : Le i j 80 | theorem leUnderSuccInv {i js} (h : LeUnder (succ i) (succOrdinals js)) : LeUnder i js 81 | theorem leOverSuccInv {is j} (h : LeOver (succOrdinals is) (succ j)) : LeOver is j 82 | by 83 | case' leSuccInv => 84 | generalize ei : succ i = si at h 85 | generalize ej : succ j = sj at h 86 | case' leUnderSuccInv => 87 | generalize ei : succ i = si at h 88 | generalize ejs : succOrdinals js = sjs at h 89 | case' leOverSuccInv => 90 | generalize eis : succOrdinals is = sis at h 91 | generalize ej : succ j = sj at h 92 | mutual_induction h, h, h 93 | case succ ih => 94 | injection ei with ei; subst ei 95 | injection ej with ej; subst ej 96 | assumption 97 | case under => injection ej 98 | case over => injection ei 99 | case underHere => 100 | subst ei; cases js 101 | case nil => unfold succOrdinals at ejs; injection ejs 102 | case cons ih _ _ => 103 | injection ejs with e₁ e₂; subst e₁; subst e₂ 104 | exact underHere (ih rfl rfl) 105 | case underThere => 106 | subst ei; cases js 107 | case nil => unfold succOrdinals at ejs; injection ejs 108 | case cons ih _ _ => 109 | injection ejs with e₁ e₂; subst e₁; subst e₂ 110 | exact underThere (ih rfl rfl) 111 | case overNil => 112 | cases is 113 | case nil => constructor 114 | case cons => injection eis 115 | case overCons => 116 | subst ej; cases is 117 | case nil => constructor 118 | case cons ihle ihleover _ _ => 119 | injection eis with e₁ e₂; subst e₁; subst e₂ 120 | exact overCons (ihle rfl rfl) (ihleover rfl rfl) 121 | 122 | joint 123 | theorem leSucc {i j} (h : Le i j) : Le i (.succ j) 124 | theorem leUnderSucc {i js} (h : LeUnder i js) : LeUnder i (succOrdinals js) 125 | theorem leOverSucc {is j} (h : LeOver is j) : LeOver is (.succ j) 126 | by 127 | mutual_induction h, h, h 128 | all_goals try simp 129 | case succ ih => exact .succ ih 130 | case under i _ _ _ => 131 | cases i 132 | case succ ih => constructor; constructor; apply leUnderSuccInv ih 133 | case limit => constructor; sorry 134 | case over => constructor; assumption 135 | case underHere => apply LeUnder.underHere; assumption 136 | case underThere => apply LeUnder.underThere; assumption 137 | case overNil => constructor 138 | case overCons => constructor; assumption; assumption 139 | -------------------------------------------------------------------------------- /MutualInductionTest/TypeTheory.lean: -------------------------------------------------------------------------------- 1 | import MutualInduction 2 | 3 | inductive Term : Type where 4 | | var : Nat → Term 5 | | type : Term 6 | | arr : Term → Term → Term 7 | | abs : Term → Term 8 | | app : Term → Term → Term 9 | 10 | inductive Ctxt : Type where 11 | | nil : Ctxt 12 | | cons : Ctxt → Term → Ctxt 13 | notation:50 "⬝" => Ctxt.nil 14 | infixl:50 "∷" => Ctxt.cons 15 | 16 | axiom subst : Term → Term → Term 17 | axiom In : Nat → Term → Ctxt → Prop 18 | axiom Eqv : Term → Term → Prop 19 | axiom refl : ∀ {a}, Eqv a a 20 | axiom trans : ∀ {a b c}, Eqv a b → Eqv b c → Eqv a c 21 | notation:40 Γ:41 "∋" x:41 "∶" A:41 => In x A Γ 22 | notation:40 a:41 "≡" b:41 => Eqv a b 23 | 24 | section 25 | set_option hygiene false 26 | local notation:40 "⊢" Γ:40 => Wf Γ 27 | local notation:40 Γ:41 "⊢" a:41 "∶" A:41 => Wt Γ a A 28 | open Term 29 | 30 | mutual 31 | -- ⊢ Γ 32 | inductive Wf : Ctxt → Prop where 33 | | nil : ⊢ ⬝ 34 | | cons {Γ A} : 35 | ⊢ Γ → 36 | Γ ⊢ A ∶ type → 37 | --------------- 38 | ⊢ Γ ∷ A 39 | 40 | -- Γ ⊢ a ∶ A 41 | inductive Wt : Ctxt → Term → Term → Prop where 42 | | var {Γ x A} : 43 | Γ ∋ x ∶ A → 44 | ⊢ Γ → 45 | --------------- 46 | Γ ⊢ var x ∶ A 47 | | type {Γ} : 48 | ⊢ Γ → 49 | ----------------- 50 | Γ ⊢ type ∶ type 51 | | arr {Γ A B} : 52 | Γ ⊢ A ∶ type → 53 | Γ ∷ A ⊢ B ∶ type → 54 | -------------------- 55 | Γ ⊢ arr A B ∶ type 56 | | abs {Γ b A B} : 57 | Γ ⊢ A ∶ type → 58 | Γ ∷ A ⊢ b ∶ B → 59 | --------------------- 60 | Γ ⊢ abs b ∶ arr A B 61 | | app {Γ b a A B} : 62 | Γ ⊢ b ∶ arr A B → 63 | Γ ⊢ a ∶ A → 64 | ------------------------ 65 | Γ ⊢ app b a ∶ subst B a 66 | | conv {Γ a A B} : 67 | Γ ⊢ a ∶ A → 68 | Γ ⊢ B ∶ type → 69 | A ≡ B → 70 | --------------- 71 | Γ ⊢ a ∶ B 72 | end 73 | end 74 | 75 | notation:40 "⊢" Γ:40 => Wf Γ 76 | notation:40 Γ:41 "⊢" a:41 "∶" A:41 => Wt Γ a A 77 | open Wf Wt 78 | 79 | /- 80 | Recursors generated for ⊢ Γ and Γ ⊢ a ∶ A: 81 | 82 | Wf.rec : 83 | ∀ {motive_1 : ∀ Γ, ⊢ Γ → Prop} 84 | {motive_2 : ∀ Γ a A, Γ ⊢ a ∶ A → Prop}, 85 | ... Wf cases ... 86 | (∀ {Γ x a A} (mem : Γ ∋ x ∶ A) (h : ⊢ Γ), 87 | motive_1 Γ h → 88 | motive_2 Γ (.var x) A (Wt.var mem h)) → 89 | ... other Wt cases ... 90 | ∀ {Γ} (h : ⊢ Γ), motive_1 Γ h 91 | 92 | Wt.rec : 93 | ∀ {motive_1 : ∀ Γ, ⊢ Γ → Prop} 94 | {motive_2 : ∀ Γ a A, Γ ⊢ a ∶ A → Prop}, 95 | ... cases ... 96 | ∀ {Γ a A} (h : Γ ⊢ a ∶ A), motive_2 Γ a A h 97 | -/ 98 | 99 | theorem wtfInd 100 | -- motives 101 | (Q : ∀ {Γ}, ⊢ Γ → Prop) 102 | (P : ∀ {Γ a A}, Γ ⊢ a ∶ A → Prop) 103 | -- Wf subgoals 104 | (nil : Q nil) 105 | (cons : ∀ {Γ A} 106 | (h : ⊢ Γ) 107 | (hA : Γ ⊢ A ∶ .type), 108 | Q h → P hA → Q (cons h hA)) 109 | -- Wt subgoals 110 | (var : ∀ {Γ x A} 111 | (mem : Γ ∋ x ∶ A) 112 | (h : ⊢ Γ), 113 | Q h → P (var mem h)) 114 | (type : ∀ {Γ} 115 | (h : ⊢ Γ), 116 | Q h → P (type h)) 117 | (arr : ∀ {Γ A B} 118 | (hA : Γ ⊢ A ∶ .type) 119 | (hB : Γ ∷ A ⊢ B ∶ .type), 120 | P hA → P hB → P (arr hA hB)) 121 | (abs : ∀ {Γ b A B} 122 | (hA : Γ ⊢ A ∶ .type) 123 | (hb : Γ ∷ A ⊢ b ∶ B), 124 | P hA → P hb → P (abs hA hb)) 125 | (app : ∀ {Γ b a A B} 126 | (hb : Γ ⊢ b ∶ .arr A B) 127 | (ha : Γ ⊢ a ∶ A), 128 | P hb → P ha → P (app hb ha)) 129 | (conv : ∀ {Γ a A B} 130 | (ha : Γ ⊢ a ∶ A) 131 | (hB : Γ ⊢ B ∶ .type) 132 | (e : A ≡ B), 133 | P ha → P hB → P (conv ha hB e)) 134 | -- goal 135 | : (∀ {Γ} (h : ⊢ Γ), Q h) ∧ (∀ {Γ a A} (h : Γ ⊢ a ∶ A), P h) := by 136 | constructor 137 | case' left => intro Γ h 138 | case' right => intro Γ a A h 139 | mutual_induction h, h 140 | case nil => exact nil 141 | case cons => apply cons <;> assumption 142 | case var => apply var <;> assumption 143 | case type => apply type <;> assumption 144 | case arr => apply arr <;> assumption 145 | case abs => apply abs <;> assumption 146 | case app => apply app <;> assumption 147 | case conv => apply conv <;> assumption 148 | 149 | def wtInd P := @Wt.rec (λ _ _ ↦ True) P (by simp) (by simp) 150 | 151 | def wfInd (P : ∀ {Γ}, ⊢ Γ → Prop) 152 | (nil : P .nil) 153 | (cons : ∀ {Γ A} (h : ⊢ Γ) (hA : Γ ⊢ A ∶ .type), P (.cons h hA)) 154 | : ∀ {Γ} (h : ⊢ Γ), P h := by 155 | apply @Wf.rec @P (λ _ _ _ _ ↦ True) nil 156 | case cons => intros; apply_rules 157 | case var | type | arr | abs | app | conv => simp 158 | 159 | theorem wtwf {Γ a A} (h : Γ ⊢ a ∶ A) : ⊢ Γ := by 160 | have _ (h : ⊢ Γ) : True := ?dummy 161 | mutual_induction h, h 162 | all_goals try simp at * <;> assumption 163 | -- induction h using Wt.rec (motive_1 := λ _ _ ↦ True) 164 | 165 | theorem regularity {Γ} : 166 | (∀ x A, Γ ∋ x ∶ A → ⊢ Γ → Γ ⊢ A ∶ .type) ∧ 167 | (∀ a A, Γ ⊢ a ∶ A → Γ ⊢ A ∶ .type) := by 168 | constructor 169 | case' left => intro x A mem h 170 | case' right => intro a A h 171 | mutual_induction h, h 172 | case nil => sorry 173 | case cons => sorry 174 | case var mem _ ih => exact ih _ _ mem 175 | case type => constructor; assumption 176 | case arr => constructor; apply wtwf; assumption 177 | case abs => constructor <;> assumption 178 | case app => sorry 179 | case conv => assumption 180 | -------------------------------------------------------------------------------- /JointTest/EvenOdd.lean: -------------------------------------------------------------------------------- 1 | import MutualInduction 2 | import Joint 3 | 4 | mutual 5 | inductive Even : Nat → Prop where 6 | | zero : Even 0 7 | | succ n : Odd n → Even (n + 1) 8 | 9 | inductive Odd : Nat → Prop where 10 | | succ n : Even n → Odd (n + 1) 11 | end 12 | open Even Odd 13 | 14 | /- 15 | Recursors generated for Even and Odd: 16 | 17 | Even.rec : 18 | ∀ {motive_1 : ∀ n, Even n → Prop} 19 | {motive_2 : ∀ n, Odd n → Prop}, 20 | motive_1 0 zero → 21 | (∀ n (on : Odd n), motive_2 n on → motive_1 (n + 1) (succ n on)) → 22 | (∀ n (en : Even n), motive_1 n en → motive_1 (n + 1) (succ n en)) → 23 | ∀ {n} (en : Even n), motive_1 n en 24 | 25 | Odd.rec : 26 | ∀ {motive_1 : ∀ n, Even n → Prop} 27 | {motive_2 : ∀ n, Odd n → Prop}, 28 | motive_1 0 zero → 29 | (∀ n (on : Odd n), motive_2 n on → motive_1 (n + 1) (succ n on)) → 30 | (∀ n (en : Even n), motive_1 n en → motive_1 (n + 1) (succ n en)) → 31 | ∀ {n} (on : Odd n), motive_2 n on 32 | -/ 33 | 34 | theorem evenOddInd 35 | -- motives 36 | (P : ∀ {n}, Even n → Prop) 37 | (Q : ∀ {n}, Odd n → Prop) 38 | -- subgoals 39 | (ezero : P zero) 40 | (esucc : ∀ n (on : Odd n), Q on → P (succ n on)) 41 | (osucc : ∀ n (en : Even n), P en → Q (succ n en)) : 42 | -- goal 43 | ∀ n, (∀ (en : Even n), P en) ∧ (∀ (on : Odd n), Q on) := by 44 | intros n; constructor 45 | case' left => intro en -- apply @Even.rec @P @Q <;> assumption 46 | case' right => intro on -- apply @Odd.rec @P @Q <;> assumption 47 | mutual_induction on, en 48 | all_goals apply_rules 49 | 50 | theorem evenOddInd' 51 | -- motives 52 | (P : ∀ {n}, Even n → Prop) 53 | (Q : ∀ {n}, Odd n → Prop) 54 | -- subgoals 55 | (ezero : P zero) 56 | (esucc : ∀ n (on : Odd n), Q on → P (succ n on)) 57 | (osucc : ∀ n (en : Even n), P en → Q (succ n en)) : 58 | -- goal 59 | ∀ n, (∀ (en : Even n), P en) ∧ (∀ (on : Odd n), Q on) := by 60 | intro n; induction n 61 | case' succ ih => let ⟨ihe, iho⟩ := ih 62 | all_goals constructor <;> intro h <;> cases h <;> apply_rules 63 | 64 | joint (n : Nat) 65 | theorem evenInv (h : Even (n + 1)) : Odd n 66 | theorem oddInv (h : Odd (n + 1)) : Even n 67 | by 68 | all_goals generalize e : n + 1 = m at h 69 | mutual_induction h, h 70 | case zero => injection e 71 | all_goals injection e with e; subst e; assumption 72 | 73 | joint {n m : Nat} 74 | /-- even sums are either both even or both odd -/ 75 | theorem plusEven (enm : Even (n + m)) : (Even n ∧ Even m) ∨ (Odd n ∧ Odd m) 76 | /-- odd sums are even in one and odd in the other -/ 77 | theorem plusOdd (onm : Odd (n + m)) : (Odd n ∧ Even m) ∨ (Even n ∧ Odd m) 78 | by 79 | case' plusOdd => generalize e₂ : n + m = k₂ at onm 80 | case' plusEven => generalize e₁ : n + m = k₁ at enm 81 | mutual_induction enm, onm generalizing n 82 | case plusEven.zero => 83 | have _ : n = 0 := by omega 84 | have _ : m = 0 := by omega 85 | subst n m 86 | left; constructor <;> constructor 87 | case plusEven.succ => 88 | cases n 89 | case zero => 90 | simp at e₁; subst e₁ 91 | left; constructor <;> constructor; assumption 92 | case succ k _ ih n => 93 | have e : n + m = k := by omega 94 | cases ih e 95 | case inl h => 96 | let ⟨en, om⟩ := h 97 | left; constructor; constructor; assumption; assumption 98 | case inr h => 99 | let ⟨en, om⟩ := h 100 | right; constructor; constructor; assumption; assumption 101 | case plusOdd.succ => 102 | cases n 103 | case zero => 104 | simp at e₂; subst e₂ 105 | right; constructor <;> constructor; assumption 106 | case succ k _ ih n => 107 | have e : n + m = k := by omega 108 | cases ih e 109 | case inl h => 110 | let ⟨en, om⟩ := h 111 | left; constructor; constructor; assumption; assumption 112 | case inr h => 113 | let ⟨en, om⟩ := h 114 | right; constructor; constructor; assumption; assumption 115 | 116 | theorem even2 n (en : Even (n + 2)) : Even n := by 117 | generalize e : n + 2 = m at en 118 | mutual_induction en generalizing n 119 | case zero => injection e 120 | case succ on => cases e; cases on; assumption 121 | 122 | mutual 123 | variable 124 | {motive_1 : ∀ n, Even n → Prop} 125 | {motive_2 : ∀ n, Odd n → Prop} 126 | 127 | inductive Even.below' : ∀ {n}, Even n → Prop where 128 | | zero : Even.below' zero 129 | | succ : ∀ n {on : Odd n}, motive_2 n on → 130 | Odd.below' on → Even.below' (succ n on) 131 | 132 | inductive Odd.below' : ∀ {n}, Odd n → Prop where 133 | | succ : ∀ n {en : Even n}, motive_1 n en → 134 | Even.below' en → Odd.below' (succ n en) 135 | end 136 | 137 | joint 138 | {P : ∀ n, Even n → Prop} {Q : ∀ n, Odd n → Prop} {n : Nat} 139 | (he : ∀ {n} {en : Even n}, @Even.below P Q n en → P n en) 140 | (ho : ∀ {n} {on : Odd n}, @Odd.below P Q n on → Q n on) 141 | theorem Even.brecOn' : ∀ (en : Even n), P n en 142 | theorem Odd.brecOn' : ∀ (on : Odd n), Q n on 143 | by 144 | case' Odd.brecOn' => intro on; apply ho 145 | case' Even.brecOn' => intro en; apply he 146 | mutual_induction en, on using Even.rec, Odd.rec 147 | case zero => exact Even.below.zero 148 | case Even.brecOn'.succ ih => exact Even.below.succ _ _ ih (ho ih) 149 | case Odd.brecOn'.succ ih => exact Odd.below.succ _ _ ih (he ih) 150 | -------------------------------------------------------------------------------- /MutualInductionTest/EvenOdd.lean: -------------------------------------------------------------------------------- 1 | import MutualInduction 2 | 3 | mutual 4 | inductive Even : Nat → Prop where 5 | | zero : Even 0 6 | | succ n : Odd n → Even (n + 1) 7 | 8 | inductive Odd : Nat → Prop where 9 | | succ n : Even n → Odd (n + 1) 10 | end 11 | open Even Odd 12 | 13 | /- 14 | Recursors generated for Even and Odd: 15 | 16 | Even.rec : 17 | ∀ {motive_1 : ∀ n, Even n → Prop} 18 | {motive_2 : ∀ n, Odd n → Prop}, 19 | motive_1 0 zero → 20 | (∀ n (on : Odd n), motive_2 n on → motive_1 (n + 1) (succ n on)) → 21 | (∀ n (en : Even n), motive_1 n en → motive_1 (n + 1) (succ n en)) → 22 | ∀ {n} (en : Even n), motive_1 n en 23 | 24 | Odd.rec : 25 | ∀ {motive_1 : ∀ n, Even n → Prop} 26 | {motive_2 : ∀ n, Odd n → Prop}, 27 | motive_1 0 zero → 28 | (∀ n (on : Odd n), motive_2 n on → motive_1 (n + 1) (succ n on)) → 29 | (∀ n (en : Even n), motive_1 n en → motive_1 (n + 1) (succ n en)) → 30 | ∀ {n} (on : Odd n), motive_2 n on 31 | -/ 32 | 33 | theorem evenOddInd 34 | -- motives 35 | (P : ∀ {n}, Even n → Prop) 36 | (Q : ∀ {n}, Odd n → Prop) 37 | -- subgoals 38 | (ezero : P zero) 39 | (esucc : ∀ n (on : Odd n), Q on → P (succ n on)) 40 | (osucc : ∀ n (en : Even n), P en → Q (succ n en)) : 41 | -- goal 42 | ∀ n, (∀ (en : Even n), P en) ∧ (∀ (on : Odd n), Q on) := by 43 | intros n; constructor 44 | case' left => intro en -- apply @Even.rec @P @Q <;> assumption 45 | case' right => intro on -- apply @Odd.rec @P @Q <;> assumption 46 | mutual_induction on, en 47 | all_goals apply_rules 48 | 49 | theorem evenOddInd' 50 | -- motives 51 | (P : ∀ {n}, Even n → Prop) 52 | (Q : ∀ {n}, Odd n → Prop) 53 | -- subgoals 54 | (ezero : P zero) 55 | (esucc : ∀ n (on : Odd n), Q on → P (succ n on)) 56 | (osucc : ∀ n (en : Even n), P en → Q (succ n en)) : 57 | -- goal 58 | ∀ n, (∀ (en : Even n), P en) ∧ (∀ (on : Odd n), Q on) := by 59 | intro n; induction n 60 | case' succ ih => let ⟨ihe, iho⟩ := ih 61 | all_goals constructor <;> intro h <;> cases h <;> apply_rules 62 | 63 | theorem evenOddInv : 64 | (∀ n m, m = n + 1 → Even m → Odd n) ∧ 65 | (∀ n m, m = n + 1 → Odd m → Even n) := by 66 | constructor 67 | all_goals intro n m e h 68 | mutual_induction h, h 69 | case zero => injection e 70 | all_goals injection e with e; subst e; assumption 71 | 72 | theorem evenOddInv' (n : Nat) : 73 | (∀ m, m = n + 1 → Even m → Odd n) ∧ 74 | (∀ m, m = n + 1 → Odd m → Even n) := by 75 | constructor 76 | all_goals intro m e h 77 | mutual_induction h, h 78 | case zero => injection e 79 | all_goals injection e with e; subst e; assumption 80 | 81 | theorem plusEvenOdd n m : 82 | (Even (n + m) → (Even n ∧ Even m) ∨ (Odd n ∧ Odd m)) ∧ 83 | (Odd (n + m) → (Odd n ∧ Even m) ∨ (Even n ∧ Odd m)) := by 84 | constructor 85 | case' right => intro onm; generalize e₂ : n + m = k₂ at onm 86 | case' left => intro enm; generalize e₁ : n + m = k₁ at enm 87 | mutual_induction enm, onm generalizing n 88 | case left.zero => 89 | have _ : n = 0 := by omega 90 | have _ : m = 0 := by omega 91 | subst n m 92 | left; constructor <;> constructor 93 | case left.succ => 94 | cases n 95 | case zero => 96 | simp at e₁; subst e₁ 97 | left; constructor <;> constructor; assumption 98 | case succ k _ ih n => 99 | have e : n + m = k := by omega 100 | cases ih n e 101 | case inl h => 102 | let ⟨en, om⟩ := h 103 | left; constructor; constructor; assumption; assumption 104 | case inr h => 105 | let ⟨en, om⟩ := h 106 | right; constructor; constructor; assumption; assumption 107 | case right.succ => 108 | cases n 109 | case zero => 110 | simp at e₂; subst e₂ 111 | right; constructor <;> constructor; assumption 112 | case succ k _ ih n => 113 | have e : n + m = k := by omega 114 | cases ih n e 115 | case inl h => 116 | let ⟨en, om⟩ := h 117 | left; constructor; constructor; assumption; assumption 118 | case inr h => 119 | let ⟨en, om⟩ := h 120 | right; constructor; constructor; assumption; assumption 121 | 122 | theorem even2 n (en : Even (n + 2)) : Even n := by 123 | generalize e : n + 2 = m at en 124 | mutual_induction en generalizing n 125 | case zero => injection e 126 | case succ on => cases e; cases on; assumption 127 | 128 | mutual 129 | variable 130 | {motive_1 : ∀ n, Even n → Prop} 131 | {motive_2 : ∀ n, Odd n → Prop} 132 | 133 | inductive Even.below' : ∀ {n}, Even n → Prop where 134 | | zero : Even.below' zero 135 | | succ : ∀ n {on : Odd n}, motive_2 n on → 136 | Odd.below' on → Even.below' (succ n on) 137 | 138 | inductive Odd.below' : ∀ {n}, Odd n → Prop where 139 | | succ : ∀ n {en : Even n}, motive_1 n en → 140 | Even.below' en → Odd.below' (succ n en) 141 | end 142 | 143 | theorem brecOn 144 | {P : ∀ n, Even n → Prop} 145 | {Q : ∀ n, Odd n → Prop} : 146 | (∀ {n} {en : Even n}, @Even.below P Q n en → P n en) → 147 | (∀ {n} {on : Odd n}, @Odd.below P Q n on → Q n on) → 148 | (∀ {n} (en : Even n), P n en) ∧ 149 | (∀ {n} (on : Odd n), Q n on) := by 150 | intro he ho 151 | constructor 152 | case' left => intro n en; apply he 153 | case' right => intro n on; apply ho 154 | mutual_induction on, en using Odd.rec, Even.rec 155 | case zero => exact Even.below.zero 156 | case left.succ ih => exact Even.below.succ _ _ ih (ho ih) 157 | case right.succ ih => exact Odd.below.succ _ _ ih (he ih) 158 | -------------------------------------------------------------------------------- /MutualInductionTest/Ordinals.lean: -------------------------------------------------------------------------------- 1 | import MutualInduction 2 | 3 | mutual 4 | inductive Ordinal : Type where 5 | | succ : Ordinal → Ordinal 6 | | limit : Ordinals → Ordinal 7 | 8 | inductive Ordinals : Type where 9 | | nil : Ordinals 10 | | cons : Ordinal → Ordinals → Ordinals 11 | end 12 | open Ordinal Ordinals 13 | 14 | mutual 15 | inductive Le : Ordinal → Ordinal → Prop where 16 | | succ {i j} : Le i j → Le (.succ i) (.succ j) 17 | | under {i js} : LeUnder i js → Le i (.limit js) 18 | | over {is j} : LeOver is j → Le (.limit is) j 19 | 20 | inductive LeUnder : Ordinal → Ordinals → Prop where 21 | | underHere {i j js} : Le i j → LeUnder i (.cons j js) 22 | | underThere {i j js} : LeUnder i js → LeUnder i (.cons j js) 23 | 24 | inductive LeOver : Ordinals → Ordinal → Prop where 25 | | overNil {j} : LeOver .nil j 26 | | overCons {i is j} : Le i j → LeOver is j → LeOver (.cons i is) j 27 | end 28 | open LeUnder LeOver 29 | 30 | def zero := Ordinal.limit .nil 31 | theorem zeroLe {i} : Le zero i := Le.over .overNil 32 | theorem noLeZero {i} : Le (.succ i) zero → False := by 33 | intro le; cases le 34 | case under leUnder => cases leUnder 35 | 36 | theorem leInd 37 | -- motives 38 | (P : ∀ {i j}, Le i j → Prop) 39 | (Q : ∀ {i js}, LeUnder i js → Prop) 40 | (R : ∀ {is j}, LeOver is j → Prop) 41 | -- Le subgoals 42 | (succ : ∀ {i j} 43 | (h : Le i j), 44 | P h → P (.succ h)) 45 | (under : ∀ {i js} 46 | (h : LeUnder i js), 47 | Q h → P (.under h)) 48 | (over : ∀ {is j} 49 | (h : LeOver is j), 50 | R h → P (.over h)) 51 | -- LeUnder subgoals 52 | (underHere : ∀ {i j js} 53 | (h : Le i j), 54 | P h → Q (.underHere (js := js) h)) 55 | (underThere : ∀ {i j js} 56 | (h : LeUnder i js), 57 | Q h → Q (.underThere (j := j) h)) 58 | -- LeOver subgoals 59 | (overNil : ∀ {j}, R (.overNil (j := j))) 60 | (overCons : ∀ {i is j} 61 | (hle : Le i j) 62 | (h : LeOver is j), 63 | P hle → R h → R (.overCons hle h)) 64 | : (∀ {i j} (h : Le i j), P h) ∧ 65 | (∀ {i js} (h : LeUnder i js), Q h) ∧ 66 | (∀ {is j} (h : LeOver is j), R h) := by 67 | repeat' constructor 68 | case' left => intro _ _ le 69 | case' right.left => intro _ _ leunder 70 | case' right.right => intro _ _ leover 71 | mutual_induction leover, leunder, le 72 | case underThere ih => apply underThere _ ih 73 | all_goals apply_rules 74 | 75 | @[simp] 76 | def succOrdinals : Ordinals → Ordinals 77 | | nil => nil 78 | | cons i is => cons (succ i) (succOrdinals is) 79 | 80 | theorem leSuccInv' : 81 | (∀ {i j}, Le (succ i) (succ j) → Le i j) ∧ 82 | (∀ {i js}, LeUnder (succ i) (succOrdinals js) → LeUnder i js) ∧ 83 | (∀ {is j}, LeOver (succOrdinals is) (succ j) → LeOver is j) := by 84 | repeat' constructor 85 | case' left => 86 | intro i j h 87 | generalize ei : succ i = si at h 88 | generalize ej : succ j = sj at h 89 | case' right.left => 90 | intro i js h 91 | generalize ei : succ i = si at h 92 | generalize ejs : succOrdinals js = sjs at h 93 | case' right.right => 94 | intro is j h 95 | generalize eis : succOrdinals is = sis at h 96 | generalize ej : succ j = sj at h 97 | mutual_induction h, h, h 98 | case succ ih => 99 | injection ei with ei; subst ei 100 | injection ej with ej; subst ej 101 | assumption 102 | case under => injection ej 103 | case over => injection ei 104 | case underHere => 105 | subst ei; cases js 106 | case nil => unfold succOrdinals at ejs; injection ejs 107 | case cons ih _ _ => 108 | injection ejs with e₁ e₂; subst e₁; subst e₂ 109 | exact underHere (ih rfl rfl) 110 | case underThere => 111 | subst ei; cases js 112 | case nil => unfold succOrdinals at ejs; injection ejs 113 | case cons ih _ _ => 114 | injection ejs with e₁ e₂; subst e₁; subst e₂ 115 | exact underThere (ih rfl rfl) 116 | case overNil => 117 | cases is 118 | case nil => constructor 119 | case cons => injection eis 120 | case overCons => 121 | subst ej; cases is 122 | case nil => constructor 123 | case cons ihle ihleover _ _ => 124 | injection eis with e₁ e₂; subst e₁; subst e₂ 125 | exact overCons (ihle rfl rfl) (ihleover rfl rfl) 126 | 127 | theorem leSuccInv : ∀ {i j}, Le (.succ i) (.succ j) → Le i j := by 128 | intro i j le 129 | let ⟨h, _, _⟩ := leSuccInv' 130 | exact h le 131 | 132 | theorem leUnderSuccInv : ∀ {i js}, LeUnder (.succ i) (succOrdinals js) → LeUnder i js := by 133 | intro i js leUnder 134 | let ⟨_, h, _⟩ := leSuccInv' 135 | exact h leUnder 136 | 137 | theorem leOverSuccInv : ∀ {is j}, LeOver (succOrdinals is) (.succ j) → LeOver is j := by 138 | intro is j leOver 139 | let ⟨_, _, h⟩ := leSuccInv' 140 | exact h leOver 141 | 142 | theorem leSucc : 143 | (∀ {i j}, Le i j → Le i (.succ j)) ∧ 144 | (∀ {i js}, LeUnder i js → LeUnder i (succOrdinals js)) ∧ 145 | (∀ {is j}, LeOver is j → LeOver is (.succ j)) := by 146 | repeat' constructor 147 | all_goals intros _ _ h 148 | mutual_induction h, h, h 149 | all_goals try simp 150 | case succ ih => exact .succ ih 151 | case under i _ _ _ => 152 | cases i 153 | case succ ih => constructor; constructor; apply leUnderSuccInv ih 154 | case limit => constructor; sorry 155 | case over => constructor; assumption 156 | case underHere => apply LeUnder.underHere; assumption 157 | case underThere => apply LeUnder.underThere; assumption 158 | case overNil => constructor 159 | case overCons => constructor; assumption; assumption 160 | -------------------------------------------------------------------------------- /Joint.lean: -------------------------------------------------------------------------------- 1 | module 2 | 3 | public meta import Lean.Parser.Term 4 | 5 | namespace Lean.Macro 6 | open Lean.Parser.Term 7 | 8 | declare_syntax_cat theoremDecl 9 | syntax declModifiers "theorem " ident ppIndent(declSig) : theoremDecl 10 | public meta def binder := bracketedBinder (requireType := true) 11 | 12 | syntax (name := joint) 13 | "joint" (".{" ident,+ "}")? binder* 14 | theoremDecl+ byTactic' : command 15 | 16 | def binderKinds : SyntaxNodeKinds := 17 | [`ident, `Lean.Parser.Term.hole, `Lean.Parser.Term.bracketedBinder] 18 | 19 | instance : Coe (TSyntax binderKinds) (TSyntax `Lean.Parser.Term.funBinder) where 20 | coe stx := ⟨stx.raw⟩ 21 | 22 | structure JointVars where 23 | univs : TSyntaxArray `ident 24 | binders : TSyntaxArray `Lean.Parser.Term.bracketedBinder 25 | 26 | structure TheoremDecl where 27 | /-- Declaration modifiers -/ 28 | mods : TSyntax `Lean.Parser.Command.declModifiers 29 | /-- Syntax object of `theorem` keyword -/ 30 | stx : Syntax 31 | name : TSyntax `ident 32 | binders : TSyntaxArray binderKinds 33 | /-- Result type of declaration -/ 34 | sig : TSyntax `term 35 | 36 | /-- Join a sequence of names by underscores, 37 | preceded and postceded by underscores. -/ 38 | meta def mkJointName (names : Array Name) : Name := 39 | match names.foldl append (Name.mkStr1 "") with 40 | | .str n s | .num n s => .str n s!"{s}_" 41 | | n => n 42 | where append : Name → Name → Name 43 | | .str n s₁, .str _ s₂ 44 | | .str n s₁, .num _ s₂ => 45 | .str n s!"{s₁}_{s₂}" 46 | | n, _ => n 47 | 48 | /-- Get the bound variable `x` in a typed bracketed binder 49 | `(x : A)`, `{x : A}`, `⦃x : A⦄`, or `{{x : A}}`. -/ 50 | meta def getBoundVars (binder : TSyntax `Lean.Parser.Term.bracketedBinder) : TSyntaxArray `ident := 51 | match binder with 52 | | `(binder| ($xs:ident* : $_:term)) 53 | | `(binder| {$xs:ident* : $_:term}) 54 | | `(binder| ⦃$xs:ident* : $_:term⦄) 55 | | `(binder| {{$xs:ident* : $_:term}}) => xs 56 | | _ => #[] 57 | 58 | /-- 59 | Given a theorem declaration of the form 60 | `theorem thm (x : A)... : B`, 61 | create the pair 62 | `⟨(∀ (x : A)..., B), (λ (x : A)... ↦ ?thm)⟩` 63 | of type `(Σ' A : Sort _, A)`. 64 | -/ 65 | meta def mkThmPair (thm : TheoremDecl) : MacroM (TSyntax `term) := do 66 | `(PSigma.mk (∀ $thm.binders*, $thm.sig) (λ $thm.binders* ↦ ?$thm.name)) 67 | 68 | /-- 69 | Given `n` theorems `thmᵢ : Aᵢ`, create a definition named `_thm₁_..._thmₙ_` 70 | with a heterogeneous array of proofs of `Aᵢ`. 71 | The body of the definition is a refined array of holes `?thmᵢ : A_i` 72 | that should then be solved by the given `tactics`. 73 | -/ 74 | meta def mkJointDef (jnt : JointVars) (thms : Array TheoremDecl) (byTk : Syntax) (tactics : Syntax.TSepArray `tactic "") : 75 | MacroM (Command × TSyntax `ident) := do 76 | let id := mkIdent <| mkJointName (thms.map (·.name.getId)) 77 | let pairs ← thms.mapM mkThmPair 78 | let byBlock ← withRef byTk `(term| by refine #[$pairs,*]; $tactics*) 79 | let defn ← 80 | if jnt.univs.isEmpty then 81 | `(command| abbrev $id $jnt.binders* : Array (PSigma fun (A : Sort _) ↦ A) := $byBlock) 82 | else 83 | `(command| abbrev $id.{$jnt.univs,*} $jnt.binders* : Array (PSigma fun (A : Sort _) ↦ A) := $byBlock) 84 | return (defn, id) 85 | 86 | /-- 87 | The `i`th theorem is proven by the `i`th element of the jointly defined array: 88 | `theorem thmₙ : Aₙ := _thm₁_..._thmₙ_[i].snd`. 89 | Note it must be that `_thm₁_..._thmₙ_[i].fst = Aₙ`. 90 | -/ 91 | meta def mkNthThm (id : TSyntax `ident) (jnt : JointVars) (i : Nat) (thm : TheoremDecl) : MacroM Command := do 92 | let istx := Syntax.mkNatLit i 93 | let args := Array.flatten <| jnt.binders.map getBoundVars 94 | let nthThm ← withRef thm.stx <| 95 | if jnt.univs.isEmpty then 96 | `(command| $thm.mods:declModifiers theorem $thm.name $jnt.binders* : ∀ $thm.binders*, $thm.sig := (@$id $args*)[$istx].snd) 97 | else 98 | `(command| $thm.mods:declModifiers theorem $thm.name.{$jnt.univs,*} $jnt.binders* : ∀ $thm.binders*, $thm.sig := (@$id.{$jnt.univs,*} $args*)[$istx].snd) 99 | `(command| set_option linter.unusedVariables false in $nthThm) 100 | 101 | /-- 102 | Given theorem statements of the form `theorem thm (y : B)... : C`, 103 | possibly sharing joint universe variables `u...` and term variables `(x : A)...`, 104 | the joint theorem declaration 105 | ``` 106 | joint.{u...} (x : A)... 107 | theorem thm (y : B)... : C 108 | ... 109 | by ... 110 | ``` 111 | enters a proof state containing each theorem as a separate goal. 112 | This is different from `mutual` theorems because their bodies cannot be mutually recursive. 113 | 114 | As an example, given the following joint theorem statements: 115 | ``` 116 | joint (n : Nat) 117 | theorem evenInv (h : isEven (n + 1)) : isOdd n 118 | theorem oddInv (h : isOdd (n + 1)) : isEven n 119 | by ... 120 | ``` 121 | the proof environment contains two goals: 122 | * `case evenInv` with state `n : Nat, h : isEven (n + 1) ⊢ isOdd n`, and 123 | * `case oddInv` with state `n : Nat, h : isOdd (n + 1) ⊢ isEven n`. 124 | -/ 125 | @[macro «joint»] 126 | public meta def expandJoint : Macro := λ stx ↦ do 127 | match stx with 128 | | `(command| joint$[.{$univs,*}]? $vars* $thms:theoremDecl* by%$byTk $tactics:tactic*) => 129 | let jointVars : JointVars := {univs := univs.getD {}, binders := vars} 130 | let thmDecls ← thms.mapM (λ (thm : TSyntax `theoremDecl) ↦ do 131 | match thm with 132 | | `(theoremDecl| $mods:declModifiers theorem%$thmTk $name:ident $binders* : $sig) => 133 | return {mods, stx := thmTk, name, binders, sig : TheoremDecl} 134 | | _ => throwUnsupported) 135 | let (jointDef, name) ← mkJointDef jointVars thmDecls byTk tactics 136 | let nthThms ← thmDecls.mapIdxM (mkNthThm name jointVars) 137 | return mkNullNode (#[jointDef] ++ nthThms) 138 | | _ => throwUnsupported 139 | 140 | end Macro 141 | -------------------------------------------------------------------------------- /MkAll.lean: -------------------------------------------------------------------------------- 1 | module 2 | 3 | import Lean 4 | 5 | namespace Lean.Meta.MkAll 6 | open Meta 7 | 8 | @[inline] 9 | def nonstrictposMsg (ctor : Name) (α : Expr) := 10 | m!"parameter {α} appears non-strictly-positively in constructor {ctor}" 11 | 12 | /-- 13 | If a free variable `fv` from `fvs` appears in the types of `exprs`, 14 | throw an error with a message about `fv`. 15 | -/ 16 | def checkFVars {ρe ρf : Type → Type} [ForIn MetaM (ρe Expr) Expr] [ForIn MetaM (ρf Expr) Expr] 17 | (exprs : ρe Expr) (fvs : ρf Expr) (emsg : Expr → MessageData): MetaM Unit := do 18 | for expr in exprs do 19 | let exprType ← inferType expr 20 | for fv in fvs do 21 | if (collectFVars {} exprType).fvarSet.contains fv.fvarId! then 22 | throwError emsg fv 23 | 24 | /-- 25 | For each parameter `α`, create a motive `α → Sort rlvl`, 26 | and `k`ontinue with the parameters and the motives. 27 | -/ 28 | def withLocalMotives {β} [Inhabited β] (αnames : Array Name) (rlvl : Level) (k : Array Expr → Array Expr → MetaM β) : MetaM β := do 29 | let mut motiveDecls : Array (Name × (Array Expr → MetaM Expr)) := #[] 30 | let mut params : Array Expr := #[] 31 | for αname in αnames do 32 | let some (.cdecl _ αfvar _ αtype _ _) := (← getLCtx).findFromUserName? αname 33 | | throwError "unknown parameter {αname}" 34 | unless αtype.isSort do 35 | throwError "the type of parameter {αname} is not a sort" 36 | let motiveType := mkForall αname .default (.fvar αfvar) (.sort rlvl) 37 | let motiveName := αname.str "motive" 38 | motiveDecls := motiveDecls.push (motiveName, λ _ ↦ pure motiveType) 39 | params := params.push (.fvar αfvar) 40 | withLocalDeclsD motiveDecls (k params) 41 | 42 | /-- 43 | Given a constructor `c` of an inductive type `I`, its type must have the following shape: 44 | ``` 45 | c : ∀ (param : paramType)... (field : fieldType)..., I param... index... 46 | ``` 47 | We wish to build the corresponding `I.all` constructor, 48 | which for the selected parameters `α...` from `param...` 49 | has as additional parameters `(motive_α : α → Prop)...`. 50 | This constructor has the following shape: 51 | ``` 52 | c.all : ∀ (param : paramType)... (field : fieldType)... (newField : newFieldType)..., 53 | I.all param... motive_α... index... (c param... field...) 54 | ``` 55 | The new fields are built from the existing fields 56 | when they return elements of `α...` or of `I`. 57 | That is, if `fieldType` has the shape `∀ (arg : argType)..., appFn appArgs...`, 58 | * When `appFn` is `α`, 59 | `newFieldType` is `∀ (arg : argType)..., motive_α (field arg...)`; 60 | * When `appFn` is `I` and `appArgs` is `param... cindex...`, 61 | `newFieldType` is `∀ (arg : argType)..., I.all param... motive_α... cindex... (field arg...)`; 62 | 63 | otherwise, no new field is added. 64 | -/ 65 | def buildCtorType (all : Name) (αnames : Array Name) (indVal : InductiveVal) (ctorInfo : ConstructorVal) : MetaM Constructor := do 66 | let type ← forallTelescope ctorInfo.type fun ctorArgs ctorInd => do 67 | withLocalMotives αnames .zero fun αs motives => do 68 | let αmotives := List.zip αs.toList motives.toList 69 | assert! ctorArgs.size == ctorInfo.numParams + ctorInfo.numFields 70 | let params : Array Expr := ctorArgs[:ctorInfo.numParams] 71 | let fields : Array Expr := ctorArgs[ctorInfo.numParams:] 72 | let ⟨_, indArgs⟩ := ctorInd.getAppFnArgs 73 | assert! indArgs.size == indVal.numParams + indVal.numIndices 74 | let indices : Array Expr := indArgs[indVal.numParams:] 75 | let appCtor := mkAppN (.const ctorInfo.name (ctorInfo.levelParams.map .param)) ctorArgs 76 | let ctorType := mkAppN allExpr (params ++ motives ++ indices ++ #[appCtor]) 77 | mkForallFVars (params ++ motives ++ fields) (← loop αmotives ctorType #[] fields.toList) 78 | pure { name := ctorInfo.name.updatePrefix all, type } 79 | where 80 | allExpr : Expr := .const all (indVal.levelParams.map .param) 81 | loop (αmotives : List (Expr × Expr)) (ctorType : Expr) (newFields : Array Expr) : List Expr → MetaM Expr 82 | | [] => mkForallFVars newFields ctorType 83 | | field :: fields => do 84 | let fieldType ← inferType field 85 | forallTelescope fieldType fun args type => do 86 | checkFVars args (αmotives.map (·.fst)) (nonstrictposMsg ctorInfo.name) 87 | let appFn := type.getAppFn 88 | let appliedField := mkAppN field args 89 | if let some αmotive := αmotives.lookup appFn then 90 | let mfieldType ← mkForallFVars args (mkApp αmotive appliedField) 91 | withLocalDeclD .anonymous mfieldType fun mfield => 92 | loop αmotives ctorType (newFields.push mfield) fields 93 | else if appFn.constName?.isEqSome indVal.name then 94 | let appArgs := type.getAppArgs 95 | let params : Array Expr := appArgs[:indVal.numParams] 96 | let indices : Array Expr := appArgs[indVal.numParams:] 97 | let motives := αmotives.map (·.snd) 98 | let allType := mkAppN allExpr (params ++ motives ++ indices ++ #[appliedField]) 99 | let rfieldType ← mkForallFVars args allType 100 | withLocalDeclD .anonymous rfieldType fun rfield => 101 | loop αmotives ctorType (newFields.push rfield) fields 102 | else loop αmotives ctorType newFields fields 103 | 104 | /-- 105 | Given an inductive predicate `I (param : paramType)... : ∀ (index : indexType)..., Prop` 106 | and selected parameters `α...` from `param...`, 107 | create an auxiliary predicate on `I` parametrized over `(motive_α : α → Prop)...`, 108 | which has the following type: 109 | ``` 110 | I.all : ∀ (param : paramType)... (motive_α : α → Prop)... (index : indexType)..., 111 | I param... index... → Prop 112 | ``` 113 | The shape of its constructors are detailed in the doc comment for `buildCtorType`. 114 | -/ 115 | def mkIAllDef (indVal : InductiveVal) (αnames : Array Name) (all : Name) : MetaM Declaration := do 116 | let indType ← forallTelescope indVal.type fun indArgs _ => do 117 | assert! indArgs.size == indVal.numParams + indVal.numIndices 118 | let params : Array Expr := indArgs[:indVal.numParams] 119 | let indices : Array Expr := indArgs[indVal.numParams:] 120 | 121 | withLocalMotives αnames .zero fun _ motives => do 122 | let indApp := mkAppN (.const indVal.name (indVal.levelParams.map .param)) indArgs 123 | let indPred ← mkArrow indApp (.sort .zero) 124 | let allIndType ← 125 | withNewBinderInfos (indArgs.map (·.fvarId!, BinderInfo.implicit)) <| 126 | mkForallFVars (params ++ motives ++ indices) indPred 127 | let allCtors ← indVal.ctors.mapM (getConstInfoCtor · >>= buildCtorType all αnames indVal) 128 | pure { name := all, type := allIndType, ctors := allCtors } 129 | pure <| Declaration.inductDecl 130 | indVal.levelParams (indVal.numParams + αnames.size) [indType] indVal.isUnsafe 131 | 132 | /-- 133 | Given a recursor for an inductive type `I` with parameters `param...` and motives `recMotive...`, 134 | along with selected parameters `α...` from `param...` and predicates `(motive_α : α → Prop)...`, 135 | we wish to build the minor premise of the given `minorType`. 136 | We know that the motives will be intantiated to constant functions returning sort `rlvl`, 137 | and the minor premise returns a product that collects 138 | recursive applications and applications of `motive_α`. 139 | That is, it has the shape `λ (field : fieldType)... ↦ prod × ... × prod`, 140 | and if `fieldType` has the shape `∀ (arg : argType)..., appFn appArgs...`, 141 | * When `appFn` is `α`, add `∀ (arg : argType)..., motive_α (field arg...)` to the product; 142 | * When `appFn` is `recMotive`, we know `appFn appArgs...` will be `Sort rlvl`, 143 | so replace `(field : fieldType)` with `(ih : ∀ (arg : argType)..., Sort rlvl)`, 144 | and add `∀ (arg : argType)..., ih arg...` to the product; 145 | 146 | otherwise, nothing is added to the product. 147 | 148 | (N.B. Although minor premises aren't constructors, 149 | we still call their arguments `field`s to distinguish them from *their own* arguments, 150 | and to avoid calling things `argArg`s.) 151 | -/ 152 | def buildMinorPremise (rlvl : Level) (αmotives : List (Expr × Expr)) (recMotives : Array Expr) 153 | (minorType : Expr) (emsg : Expr → MessageData) : MetaM Expr := do 154 | forallTelescope minorType fun fields _ => do loop #[] #[] fields.toList 155 | where 156 | loop (newFields : Array Expr) (prods : Array Expr) : List Expr → MetaM Expr 157 | | [] => do mkLambdaFVars newFields (← PProdN.pack rlvl prods) 158 | | field :: fields => do 159 | let fieldType ← inferType field 160 | forallTelescope fieldType fun args type => do 161 | checkFVars args (αmotives.map (·.snd)) emsg 162 | let appFn := type.getAppFn 163 | if let some motive := αmotives.lookup appFn then 164 | let mprod ← mkForallFVars args (mkApp motive (mkAppN field args)) 165 | loop (newFields.push field) (prods.push mprod) fields 166 | else if recMotives.contains appFn then 167 | let ihName ← field.fvarId!.getUserName 168 | let ihType ← mkForallFVars args (.sort rlvl) 169 | withLocalDeclD ihName ihType fun ih => do 170 | let rprod ← mkForallFVars args (mkAppN ih args) 171 | loop (newFields.push ih) (prods.push rprod) fields 172 | else loop (newFields.push field) prods fields 173 | 174 | /-- 175 | Given an inductive type `I (param : paramType)... : ∀ (index : indexType)..., Type` 176 | and selected parameters `α...` from `param...`, 177 | create an auxiliary definition over `I` parametrized over `(motive_α : α → Sort)...`, 178 | which has the following type: 179 | ``` 180 | I.all : ∀ (param : paramType)... (motive_α : α → Prop)... (index : indexType)..., 181 | I param... index... → Sort 182 | ``` 183 | To be able to use motives in arbitrary sorts, 184 | we take advantage of large elimination and build `I.all` using the recursor `I.rec`, 185 | which has the following shape: 186 | ``` 187 | I.rec : ∀ (param : paramType)... 188 | (recMotive : ∀ (index : indexType)..., I param... index... → Sort)... 189 | (minor : minorType)... 190 | (index : indexType)..., 191 | (major : I param... index...), 192 | recMotive index... major 193 | ``` 194 | Therefore, we keep `param...`, `index...`, and `major`, 195 | while we need to instantiate `recMotive...` and `minor...`. 196 | The recursor's motives are simply constant functions returning `Sort`. 197 | The shape of minor premises are detailed in the doc comment for `buildMinorPremise`. 198 | -/ 199 | def mkAllDef (indVal : InductiveVal) (recVal : RecursorVal) (αnames : Array Name) (all : Name) (prop? : Bool) : MetaM Declaration := do 200 | let allLvlParams@(lvlParam :: lvlParams) := recVal.levelParams 201 | | throwError "recursor {recVal.name} has no levelParams" 202 | let lvls : List Level := lvlParams.map .param 203 | let allLvlParams := if prop? then lvlParams else allLvlParams 204 | let rlvl : Level := if prop? then 0 else .succ (.param lvlParam) 205 | let recType := recVal.type.instantiateLevelParams [lvlParam] [rlvl] 206 | 207 | let nParams := indVal.numParams 208 | let decl ← forallTelescope recType fun refArgs _ => do 209 | assert! refArgs.size > nParams + recVal.numMotives + recVal.numMinors 210 | let params : Array Expr := refArgs[:nParams] 211 | let rmotives : Array Expr := refArgs[nParams:nParams + recVal.numMotives] 212 | let minors : Array Expr := refArgs[nParams + recVal.numMotives:nParams + recVal.numMotives + recVal.numMinors] 213 | let indices : Array Expr := refArgs[nParams + recVal.numMotives + recVal.numMinors:refArgs.size - 1] 214 | let major : Expr := refArgs[refArgs.size - 1]! 215 | 216 | withLocalMotives αnames rlvl fun αs motives => do 217 | let αmotives := List.zip αs.toList motives.toList 218 | 219 | let mut sorts := #[] 220 | for motive in rmotives do 221 | let sort ← forallTelescope (← inferType motive) fun targs _ => 222 | mkLambdaFVars targs (.sort rlvl) 223 | sorts := sorts.push sort 224 | 225 | let mut premises := #[] 226 | for minor in minors, ctor in indVal.ctors do 227 | let premise ← buildMinorPremise rlvl αmotives rmotives (← inferType minor) (nonstrictposMsg ctor) 228 | premises := premises.push premise 229 | 230 | let recFn := .const recVal.name (rlvl.succ :: lvls) 231 | let recApp := mkAppN recFn (params ++ sorts ++ premises ++ indices ++ #[major]) 232 | let allParams := params ++ motives ++ indices ++ #[major] 233 | let allType ← mkForallFVars allParams (.sort rlvl) 234 | let allExpr ← mkLambdaFVars allParams recApp 235 | mkDefinitionValInferringUnsafe all allLvlParams allType allExpr .abbrev 236 | pure <| .defnDecl decl 237 | 238 | def mkAllInductive (ind : Name) (params : Array Name) : MetaM Unit := do 239 | let indVal ← getConstInfoInduct ind 240 | let allName := Name.str ind "all" 241 | let iallName := Name.str ind "iall" 242 | if (← isInductivePredicate ind) then 243 | addDecl <| ← mkIAllDef indVal params allName 244 | mkAllAttributes allName 245 | else 246 | let .recInfo recVal ← getConstInfo (mkRecName ind) | return 247 | addDecl <| ← mkAllDef indVal recVal params allName false 248 | addDecl <| ← mkAllDef indVal recVal params iallName true 249 | mkAllAttributes allName 250 | mkAllAttributes iallName 251 | where 252 | /-- The .all and .iall predicates are reducible, protected, and simpable. -/ 253 | mkAllAttributes (all : Name) : MetaM Unit := do 254 | setReducibleAttribute all 255 | modifyEnv fun env => addProtected env all 256 | enableRealizationsForConst all 257 | 258 | syntax (name := mkAll) "mk_all" (ppSpace ident)+ : attr 259 | 260 | initialize Lean.registerBuiltinAttribute { 261 | name := `mkAll 262 | descr := "Generate an `all` predicate for an inductive with a strictly positive parameter." 263 | add := fun decl stx _ => MetaM.run' do 264 | let ⟨decl, params⟩ ← match stx with 265 | | `(attr| mk_all $[$params:ident]*) => pure (decl, params) 266 | | _ => throwError "unrecognized syntax" 267 | mkAllInductive decl (params.map (·.getId)) 268 | } 269 | 270 | end MkAll 271 | -------------------------------------------------------------------------------- /MutualInduction.lean: -------------------------------------------------------------------------------- 1 | module 2 | 3 | import all Init.Data.Array.QSort.Basic 4 | public meta import Lean.Meta.GeneralizeVars 5 | public meta import Lean.Meta.Tactic.Generalize 6 | public meta import Lean.Elab.Tactic.Induction 7 | 8 | -- Why is this not in Lean.Expr? 9 | meta def Lean.Expr.mvarId? : Expr → Option MVarId 10 | | .mvar n => some n 11 | | _ => none 12 | 13 | -- And why is this not in Init.Data.Vector.Basic? 14 | @[inline] meta def Vector.zip3 {α β γ n} (as : Vector α n) (bs : Vector β n) (cs : Vector γ n): Vector (α × β × γ) n := 15 | ⟨as.toArray.zip (bs.toArray.zip cs.toArray), by simp⟩ 16 | 17 | -- And why is this not top-level in Init.Data.Array.QSort.Basic? 18 | meta def Vector.qsort {α n} (as : Vector α n) (lt : α → α → Bool) : Vector α n := 19 | match n with 20 | | 0 => as 21 | | n + 1 => Array.qsort.sort lt as 0 n 22 | 23 | /-- 24 | Take the first `n` elements of a list, 25 | returning a vector of length `n` 26 | only if there were `n` elements to take. -/ 27 | meta def List.takeToVector? {α} (as : List α) (n : Nat) : Option (Vector α n) := 28 | let as := (as.take n).toArray 29 | if e : as.size = n then some ⟨as, e⟩ else none 30 | 31 | namespace Lean.Elab.Tactic 32 | open Meta 33 | 34 | structure Goal where 35 | /-- Syntax object of the target -/ 36 | stx : Syntax 37 | name : Name 38 | mvarId : MVarId 39 | /-- User-facing name of the target -/ 40 | target : Name 41 | /-- Target and its indices -/ 42 | targets : Array Expr 43 | genFVars : Array FVarId 44 | indVal : InductiveVal 45 | elimInfo : ElimInfo 46 | -- deriving Inhabited 47 | 48 | structure GoalWithMotives extends Goal where 49 | motive : Expr 50 | motives : Array Expr := #[] 51 | 52 | meta structure Alt extends ElimApp.Alt where 53 | numGenFVars : Nat := 0 54 | trivial : Bool := false 55 | deriving Inhabited 56 | 57 | /-- 58 | Given `n` goals to be solved, each with a variable `x₁ ... xₙ` respectively 59 | which each belong to a different mutual inductive type, 60 | `mutual_induction x₁, ..., xₙ` 61 | applies mutual induction on each `x` to each goal. 62 | It produces one goal for each constructor of the mutual inductive types, 63 | in which the target is replaced by a general instance of that constructor 64 | and an inductive hypothesis is added for each mutually recursive argument to the constructor. 65 | Note that exactly one goal and target must be provided for each mutual inductive type. 66 | 67 | * `mutual_induction x₁, ..., xₙ using r₁, ..., rₙ` 68 | allows specifying the principles of induction for each of the `n` goals, respectively. 69 | There must be exactly as many eliminators provided as there are goals targeted. 70 | Here, each `rᵢ` must be a term whose result type has the form `Cᵢ tᵢ...`, 71 | where `Cᵢ` is a bound variable representing the motive of the `i`th goal 72 | and `tᵢ...` is a (possibly empty) sequence of bound variables. 73 | Furthermore, every `rᵢ` must quantify over 74 | exactly the same motives `C₁ ... Cₙ` in exactly the same order. 75 | * `mutual_induction x₁, ..., xₙ generalizing z₁ ... zₘ`, 76 | where `z₁ ... zₘ` are variables in the contexts of all `n` goals, 77 | generalizes over `z₁ ... zₘ` before applying the induction, 78 | but then reintroduces them in each new goal produced. 79 | The net effect is that each inductive hypothesis is generalized. 80 | 81 | As an example, suppose we have mutual inductive types `Even` and `Odd` with constructors 82 | * `Even.zero : Even`, 83 | * `Even.succ : Odd → Even`, and 84 | * `Odd.succ : Even → Odd`. 85 | 86 | Given a goal `even` with hypothesis `e : Even` and type `P e`, 87 | and a goal `odd` with hypothesis `o : Odd` and type `Q o`, 88 | `mutual_induction e, o` produces three goals 89 | (where names `o`, `e`, `ih` are chosen automatically and not accessible): 90 | * `case even.zero` with state `⊢ P zero`, 91 | * `case even.succ` with state `o : Odd, ih : Q o ⊢ P (Even.succ o)`, and 92 | * `case odd.succ` with state `e : Even, ih : P e ⊢ Q (Odd.succ o)`. 93 | -/ 94 | syntax (name := mutual_induction) 95 | "mutual_induction " term,+ 96 | (" using" term,+)? 97 | (" generalizing" (ppSpace colGt term:max)+)? 98 | : tactic 99 | 100 | /-- 101 | Find a metavariable whose name is (a suffix or prefix of) `tag`, 102 | and throw an error if none exists. 103 | This is adapted from `Lean.Elab.Tactic.findTag?`. 104 | -/ 105 | meta def findTag (mvarIds : List MVarId) (tag : Name) : MetaM MVarId := do 106 | match (← mvarIds.findM? fun mvarId => return tag == (← mvarId.getDecl).userName) with 107 | | some mvarId => return mvarId 108 | | none => 109 | match (← mvarIds.findM? fun mvarId => return tag.isSuffixOf (← mvarId.getDecl).userName) with 110 | | some mvarId => return mvarId 111 | | none => 112 | match (← mvarIds.findM? fun mvarId => return tag.isPrefixOf (← mvarId.getDecl).userName) with 113 | | some mvarId => return mvarId 114 | | none => throwError m!"goal '{tag}' not found" 115 | 116 | /-- 117 | If custom arguments are provided, 118 | ensure there are exactly as many as there are targets, 119 | returning a list of optional arguments the length of the number of targets. 120 | -/ 121 | meta def countArgs {α} (mvarId : MVarId) (argName : Name) (numTargets : Nat) (args? : Option (Array α)) : MetaM (Vector (Option α) numTargets) := do 122 | let some args := args? 123 | | return Vector.replicate numTargets none 124 | if e : args.size = numTargets then 125 | return by rw [← e]; exact args.toVector.map some 126 | else 127 | throwTacticEx `mutual_induction mvarId 128 | m!"incorrect number of {argName} provided: \ 129 | expected {numTargets}, but got {args.size}" 130 | 131 | /-- 132 | Given a list of `targets`, find their transitive dependencies in the local context, 133 | as well as the dependencies of a given expression `e`. 134 | More precisely, we compute: 135 | * `fvarIds`, the free variables whose declarations depend on `targets`, 136 | but excluding `targets` themselves; and 137 | * `fvarIdDeps`, the free variables depended upon by the declarations of `fvarIds` and by `e`, 138 | again excluding `targets` themselves. 139 | 140 | For example, given the context and goal: 141 | ``` 142 | P : ∀ n, Fin n → Prop 143 | n m : Nat 144 | q : n = m 145 | f : Fin n 146 | p : P n f 147 | ------------- 148 | ⊢ P m (q ▸ f) 149 | ``` 150 | If `f` and `n` are the targets, then `fvarIds` contains `p` and `q`, 151 | and `fvarIdDeps` contains `m`, since `q` depends on it. 152 | If `P m (q ▸ f)` is the additional expression, 153 | then `fvarIdDeps` would also contain `P`. 154 | 155 | Precondition: `targets` must be free variables. 156 | Postcondition: The resulting array is sorted in order of declaration. 157 | -/ 158 | meta def getFVarsToGeneralize (targets : Array Expr) (e : Expr) : MetaM (Array FVarId) := do 159 | let targetFVars := targets.map (·.fvarId!) 160 | let fvarIds ← Meta.getFVarsToGeneralize targets 161 | let mut s := collectFVars {} (← instantiateMVars e) 162 | for fvarId in fvarIds do 163 | let decl ← fvarId.getDecl 164 | s := collectFVars s (← instantiateMVars decl.type) 165 | if let some val := decl.value? then 166 | s := collectFVars s (← instantiateMVars val) 167 | let fvarIdDeps := s.fvarIds.filter (not ∘ targetFVars.contains) 168 | return fvarIds ++ fvarIdDeps 169 | 170 | /-- 171 | Given all goals of the mutual induction, check that they exactly cover the inductive types: 172 | * The targets must all belong to inductive types in the same mutual declaration; 173 | * The targets must each belong to a different inductive type; and 174 | * [Deprecated] The targets must belong to all of the mutual inductive types. 175 | 176 | Precondition: There must exist at least one goal. 177 | -/ 178 | meta def checkTargets {n} (goals : Vector Goal n) : MetaM Unit := do 179 | for hi : i in [0:n] do 180 | for hj : j in [0:i] do 181 | let lt : j < n := Nat.lt_trans hj.right.left hi.right.left 182 | let goal1 := goals[j]'lt 183 | let goal2 := goals[i] 184 | unless goal1.indVal.all.contains goal2.indVal.name do 185 | throwTacticEx `mutual_induction goal1.mvarId 186 | m!"targets do not belong to mutual inductive types: \ 187 | {goal1.target} is in {goal1.indVal.name}, \ 188 | while {goal2.target} is in {goal2.indVal.name}" 189 | if goal1.indVal.name == goal2.indVal.name then 190 | throwTacticEx `mutual_induction goal1.mvarId 191 | m!"duplicate target inductive types: \ 192 | {goal1.target} and {goal2.target} are both in {goal1.indVal.name}" 193 | /- No longer requiring coverage of all inductives -/ 194 | -- let allIndNames := goals[0]!.indVal.all 195 | -- let targetIndNames := goals.map (·.indVal.name) 196 | -- let mut missingIndNames : Array Name := #[] 197 | -- for indName in allIndNames do 198 | -- unless targetIndNames.contains indName do 199 | -- missingIndNames := missingIndNames.push indName 200 | -- unless missingIndNames.isEmpty do 201 | -- throwTacticEx `mutual_induction goals[0]!.mvarId 202 | -- m!"missing targets for mutual inductive types: {missingIndNames}" 203 | 204 | /-- 205 | Ensure that all given "free" variables are declared in the contexts of each goal. 206 | (They're not really "free" if they're bound in the context, are they?) 207 | -/ 208 | meta def checkFVars {n} (goals : Vector Goal n) (idents : Array Syntax) : TermElabM Unit := do 209 | for goal in goals do 210 | for ident in idents do 211 | goal.mvarId.withContext do 212 | if let none ← Term.resolveId? ident then 213 | throwTacticEx `mutual_induction goal.mvarId 214 | m!"unknown identifier '{ident}' in goal '{goal.name}'" 215 | 216 | /-- 217 | Compute all motives for all goals, 218 | in each goal abstracting over the goal's targets 219 | and generalizing over the goal's free variables, 220 | unless the variable is free in all goals, 221 | preventing unnecessary generalization, 222 | except when explicitly generalized by the user. 223 | 224 | For example, given the following two goals: 225 | ``` 226 | case goal1 227 | n m : Nat 228 | e : n = m 229 | en : Even n 230 | ----------- 231 | ⊢ Even m 232 | 233 | case goal2 234 | n m : Nat 235 | e : n = m 236 | on : Odd n 237 | ---------- 238 | ⊢ Odd m 239 | ``` 240 | If `m` is shared between the two goals, then the motives are: 241 | ``` 242 | motive_1 : λ n (en : Even n) ↦ n = m → Even m 243 | motive_2 : λ n (on : Odd n) ↦ n = m → Odd m 244 | ``` 245 | But if they are not shared, or if `m` is explicitly generalized, then the motives are: 246 | ``` 247 | motive_1 : λ n (en : Even n) ↦ ∀ m, n = m → Even m 248 | motive_2 : λ n (on : Odd n) ↦ ∀ m, n = m → Odd m 249 | ``` 250 | 251 | Preconditions: 252 | * The motives that each goal's eliminator quantifies over are exactly in the same order. 253 | 254 | Postconditions: 255 | * Each of the `n` goals has `n - 1` motives; 256 | * The motives in each goal are sorted by the order its eliminator quantifies over them; and 257 | * The goals themselves are sorted by their motives in the same order. 258 | -/ 259 | meta def addMotives {n} (gs : Vector Goal n) (userFVars : Array FVarId): MetaM (Vector GoalWithMotives n) := do 260 | let gs ← gs.mapM filterGenFVars 261 | let gs ← gs.mapM genUserFVars 262 | let gs ← gs.mapM addMotive 263 | let gs := gs.qsort (·.elimInfo.motivePos < ·.elimInfo.motivePos) 264 | return gs.map (addMotives gs) 265 | where 266 | addMotives (goals : Vector GoalWithMotives n) (goal : GoalWithMotives) : GoalWithMotives := 267 | let goals := goals.toArray.filter (·.elimInfo.motivePos != goal.elimInfo.motivePos) 268 | {goal with 269 | motives := goals.map (·.motive) 270 | elimInfo := {goal.elimInfo with 271 | targetsPos := goals.map (·.elimInfo.motivePos) 272 | ++ goal.elimInfo.targetsPos}} 273 | addMotive (g : Goal) : MetaM GoalWithMotives := 274 | g.mvarId.withContext do 275 | let ⟨genFVars, goal⟩ ← sortFVarIds g.genFVars >>= g.mvarId.revert 276 | goal.withContext do 277 | let goalType ← MetavarDecl.type <$> goal.getDecl 278 | let motive ← mkLambdaFVars g.targets goalType 279 | return {g with mvarId := goal, genFVars, motive} 280 | genUserFVars (g : Goal) : MetaM Goal := 281 | g.mvarId.withContext do 282 | let forbidden ← mkGeneralizationForbiddenSet g.targets 283 | for userFVarId in userFVars do 284 | if forbidden.contains userFVarId then 285 | throwError "variable cannot be generalized because target depends on it{indentExpr (mkFVar userFVarId)}" 286 | if g.genFVars.contains userFVarId then 287 | throwError "unnecessary 'generalizing' argument, variable '{mkFVar userFVarId}' is generalized automatically" 288 | pure {g with genFVars := userFVars ++ g.genFVars} 289 | filterGenFVars (g : Goal) : MetaM Goal := do 290 | let genFVars ← g.genFVars.filterM notFreeInAnyGoal 291 | return {g with genFVars} 292 | notFreeInAnyGoal (fvarId : FVarId) : MetaM Bool := 293 | gs.anyM (notFreeInGoal fvarId) 294 | notFreeInGoal (fvarId : FVarId) (g : Goal) : MetaM Bool := 295 | g.mvarId.withContext do 296 | let lctx ← getLCtx 297 | return !lctx.contains fvarId 298 | 299 | /-- 300 | Each application of mutual recursors produces exactly the same new subgoals 301 | each corresponding to a constructor of one of the mutual inductive types; 302 | `alts` contains these subgoals for each recursor application. 303 | Therefore, only one sequence of subgoals needs to be solved, 304 | and every other sequence can be pointwise assigned the same solution. 305 | The sequence we pick out should be the subgoals that prove the motive of its recursor, 306 | since it informs which parent case name it's associated with. 307 | The case names are in an association list `tags` that map from the inductive types' names, 308 | defaulting to the provided main goal's `tag`. 309 | 310 | Preconditions: 311 | * There must exist at least one sequence of subgoals; and 312 | * All sequences of subgoals must have the same length and pointwise have the same type. 313 | -/ 314 | meta def deduplicate {n} (tag : Name) (tags : Vector (Name × Name) n) (alts : Vector (Array Alt) n) : MetaM (Array Alt) := do 315 | let mut deduped := alts[0]! 316 | -- find the canonical alternatives that prove the motive 317 | for alt in alts do 318 | for h : i in [0:alt.size] do 319 | let alti := alt[i]'h.right.left 320 | if alti.info.provesMotive then 321 | deduped := deduped.set! i alti 322 | -- assign identical alternatives to canonical alternative 323 | for alt in alts do 324 | for h : i in [0:alt.size] do 325 | let deAlt := deduped[i]!.mvarId 326 | let otherAlt := (alt[i]'h.right.left).mvarId 327 | unless deAlt == otherAlt do 328 | let altExpectedType ← inferType (.mvar deAlt) 329 | let altType ← inferType (.mvar otherAlt) 330 | if ← isDefEqGuarded altExpectedType altType then 331 | otherAlt.assign (.mvar deAlt) 332 | -- ensure root of user-facing name corresponds to the original subgoal name 333 | for alt in deduped do 334 | let .str ind cstr := alt.info.declName?.getD (← alt.mvarId.getTag) | continue 335 | let tag := (tags.toList.lookup ind).getD tag 336 | alt.mvarId.setTag <| .str tag cstr 337 | return deduped 338 | 339 | /-- 340 | Given a goal and a target in that goal, 341 | produce all the information we can glean without considering the other mutual goals: 342 | the target and its indices, its inductive type, 343 | the goal and its the free variables to generalize, 344 | and information about the eliminator to be applied. 345 | -/ 346 | meta def getSubgoal (stxgoal : TSyntax `term × Option (TSyntax `term) × MVarId) : TacticM Goal := 347 | let ⟨targetName, elim?, goal⟩ := stxgoal 348 | goal.withContext do 349 | let target ← elabTerm targetName none 350 | let indVal ← getInductiveVal goal target 351 | let elimInfo ← do 352 | let some elim := elim? 353 | | getElimInfo (mkRecName indVal.name) indVal.name 354 | getElimExprInfo (← elabTermForElim elim) indVal.name 355 | let ⟨goal, target⟩ ← generalizeTarget goal target 356 | goal.withContext do 357 | let targetUserName ← target.fvarId!.getUserName 358 | let targets ← addImplicitTargets elimInfo #[target] 359 | checkInductionTargets targets 360 | let goalType ← MetavarDecl.type <$> goal.getDecl 361 | let genFVars ← getFVarsToGeneralize targets goalType 362 | return ⟨targetName, (← goal.getDecl).userName, goal, targetUserName, targets, genFVars, indVal, elimInfo⟩ 363 | where 364 | /-- 365 | Adapted from `Lean.Elab.Tactic.getInductiveValFromMajor`, 366 | which for some reason works in the context of the main goal, 367 | not in the current context, so using it directly would not find the target, 368 | which only exists in the context of the subgoal. 369 | -/ 370 | getInductiveVal (mvarId : MVarId) (target : Expr) : MetaM InductiveVal := 371 | mvarId.withContext do 372 | let targetType ← inferType target 373 | let targetType ← whnf targetType 374 | matchConstInduct targetType.getAppFn 375 | (fun _ => throwTacticEx `mutual_induction mvarId 376 | m!"target is not an inductive type{indentExpr targetType}") 377 | (fun val _ => pure val) 378 | /-- 379 | Copied from `Lean.Elab.Tactic.elabTermForElim`, which is private. 380 | -/ 381 | elabTermForElim (stx : Syntax) : TermElabM Expr := do 382 | if stx.isIdent then 383 | if let some e ← Term.resolveId? stx (withInfo := true) then 384 | return e 385 | Term.withoutErrToSorry <| Term.withoutHeedElabAsElim do 386 | let e ← Term.elabTerm stx none (implicitLambda := false) 387 | Term.synthesizeSyntheticMVars (postpone := .no) (ignoreStuckTC := true) 388 | let e ← instantiateMVars e 389 | let e := e.eta 390 | if e.hasMVar then 391 | let r ← abstractMVars (levels := false) e 392 | return r.expr 393 | else 394 | return e 395 | /-- 396 | Adapted from `Lean.Elab.Tactic.generalizeTargets`, 397 | which for some reason also works in the context of the main goal. 398 | -/ 399 | generalizeTarget (mvarId : MVarId) (target : Expr) : TacticM (MVarId × Expr) := do 400 | mvarId.withContext do 401 | let generalize? ← do 402 | if target.isFVar then target.fvarId!.isLetVar else pure true 403 | if generalize? then 404 | let ⟨#[target], mvarId⟩ ← mvarId.generalize #[{ expr := target }] 405 | | throwTacticEx `mutual_induction mvarId 406 | m!"failed to generalize target{indentExpr target}" 407 | return ⟨mvarId, .fvar target⟩ 408 | else return ⟨mvarId, target⟩ 409 | 410 | /-- 411 | Apply the eliminator in a goal `g` and return the new metavariables to solve, 412 | each corresponding to the constructor cases of the eliminator. 413 | The new metavariables are *not* yet added to the list of goals. 414 | The motives from the other mutual goals are considered as targets. 415 | If a motive is missing, assign it the trivial predicate returning `PUnit`. 416 | 417 | Preconditions: 418 | * `g.motives` contains the other motives in the order its eliminator quantifies over them; 419 | * `g.elimInfo.elimExpr` is an eliminator that proves the motive `g.motive`; and 420 | * `g.elimInfo.targetsPos` also contains the positions of the other motives in the eliminator. 421 | -/ 422 | meta def evalSubgoal (g : GoalWithMotives) : TacticM (Array Alt) := 423 | g.mvarId.withContext do 424 | -- instantiate eliminator 425 | let result ← withRef g.stx do 426 | ElimApp.mkElimApp g.elimInfo (g.motives ++ g.targets) g.mvarId.name 427 | -- assign current motive 428 | let motiverInferredType ← inferType g.motive 429 | let motiveType ← inferType (.mvar result.motive) 430 | unless (← isDefEqGuarded motiverInferredType motiveType) do 431 | throwTacticEx `mutual_induction g.mvarId 432 | m!"type mismatch when assigning motive{indentExpr g.motive}\n\ 433 | {← mkHasTypeButIsExpectedMsg motiverInferredType motiveType}" 434 | result.motive.assign g.motive 435 | -- apply eliminator 436 | let newGoals := result.alts.map (·.mvarId) ++ result.others 437 | let elimApp ← instantiateMVars result.elimApp 438 | let unsolvedMotives ← appendMotives elimApp.getAppArgs newGoals 439 | g.mvarId.assign elimApp 440 | -- return subgoals 441 | let targetFVars := g.targets.map (·.fvarId!) 442 | let alts ← result.alts.mapM (clearTargets targetFVars) 443 | appendGoals result.others.toList 444 | alts.mapM (addNumGenFVars unsolvedMotives) 445 | where 446 | addNumGenFVars (unsolvedMotives : Array MVarId) (alt : ElimApp.Alt) : TacticM Alt := do 447 | let altType ← inferType (.mvar alt.mvarId) 448 | let altHead := altType.getForallBody.getAppFn 449 | let trivial := 450 | if let some altHeadMVar := altHead.mvarId? then 451 | unsolvedMotives.contains altHeadMVar 452 | else false 453 | if alt.info.provesMotive 454 | then return {alt with trivial, numGenFVars := g.genFVars.size} 455 | else return {alt with trivial} 456 | appendMotives (es : Array Expr) (altMVarIds : Array MVarId) : TacticM (Array MVarId) := do 457 | let mut unsolvedMotives : Array MVarId := #[] 458 | for e in es, 459 | name in g.elimInfo.elimType.getForallBinderNames do 460 | let some mvarId := e.mvarId? | continue 461 | unless altMVarIds.contains mvarId do 462 | mvarId.setTag name 463 | unsolvedMotives := unsolvedMotives.push mvarId 464 | -- add motives as goals instead of marking them as trivial 465 | /- pushGoals unsolvedMotives.toList 466 | return #[] -/ 467 | unsolvedMotives.forM trivialMotive 468 | return unsolvedMotives 469 | trivialMotive (mvarId : MVarId) : TacticM Unit := do 470 | let motiveType ← inferType (.mvar mvarId) 471 | forallTelescope motiveType fun args bodyType => do 472 | let trivial ← mkLambdaFVars args (.const `PUnit [bodyType.sortLevel!]) 473 | mvarId.assign trivial 474 | clearTargets (mvarIds : Array FVarId) (alt : ElimApp.Alt) : TacticM ElimApp.Alt := do 475 | let mvarId ← alt.mvarId.tryClearMany mvarIds 476 | return {alt with mvarId} 477 | 478 | /-- 479 | Intro the induction hypotheses of the goal, clearing away useless `PUnit`s, 480 | along with generalized variables. 481 | If the goal itself is trivial because its motive was instantiated as `PUnit`, 482 | solve the goal by `constructor`. 483 | Otherwise, add the goal to the current list of goals. 484 | -/ 485 | meta def addSubgoal (subgoal : Alt) : TacticM Unit := do 486 | let ⟨ihs, mvarId⟩ ← subgoal.mvarId.introN subgoal.info.numFields 487 | let ⟨_, mvarId⟩ ← mvarId.introNP subgoal.numGenFVars 488 | let punits ← mvarId.withContext <| ihs.filterM isPUnit 489 | let mvarId ← mvarId.tryClearMany punits 490 | if subgoal.trivial then 491 | let mvarIds ← mvarId.constructor 492 | unless mvarIds.isEmpty do 493 | throwTacticEx `mutual_induction mvarId 494 | m!"could not solve generated subgoal {subgoal.name}" 495 | else pushGoal mvarId 496 | where 497 | isPUnit (ih : FVarId) : TacticM Bool := do 498 | let ihType ← ih.getType >>= instantiateMVars 499 | let ihType ← whnf ihType 500 | return ihType.isConstOf `PUnit 501 | 502 | /-- 503 | When evaluating the mutual induction tactic, 504 | the eliminators of each goal are computed independently first. 505 | Next, the motives are computed in tandem, 506 | since each one needs to be generalized over free variables missing in the others. 507 | Only then knowing the motives for each goal can we specialize 508 | each eliminator with all motives and apply it to their goals. 509 | The mutual eliminators produce identical subgoals, 510 | which we deduplicate so that the user ony needs to solve one set of them. 511 | -/ 512 | @[tactic mutual_induction] 513 | public meta def evalMutualInduction : Tactic := λ stx ↦ do 514 | let tag ← getMainGoal >>= (·.getTag) 515 | let ⟨targetNames, elims, genFVarNames⟩ ← parse stx 516 | let elims ← countArgs (← getMainGoal) `eliminators targetNames.size elims 517 | let goals ← takeGoals targetNames.size 518 | let stxgoals := Vector.zip3 targetNames.toVector elims goals 519 | let subgoals ← stxgoals.mapM getSubgoal 520 | checkTargets subgoals 521 | checkFVars subgoals genFVarNames 522 | let genFVars ← getFVarIds genFVarNames 523 | let tags := subgoals.map (λ goal ↦ ⟨goal.indVal.name, goal.name⟩) 524 | let subgoals ← addMotives subgoals genFVars 525 | let subgoals ← subgoals.mapM evalSubgoal 526 | let subgoals ← deduplicate tag tags subgoals 527 | for subgoal in subgoals.reverse do addSubgoal subgoal 528 | where 529 | parse (stx : Syntax) := do 530 | match stx with 531 | | `(tactic| mutual_induction $targets,* $[using $elims,*]? $[generalizing $genVars*]?) => 532 | return (targets.getElems, elims.map Syntax.TSepArray.getElems, genVars.getD #[]) 533 | | _ => throwErrorAt stx "could not parse mutual_induction tactic" 534 | takeGoals (n : Nat) : TacticM (Vector MVarId n) := do 535 | let goals : List MVarId ← getUnsolvedGoals 536 | let some goals := goals.takeToVector? n 537 | | throwError m!"insufficient number of goals available: \ 538 | expected {n}, but found {goals.length}" 539 | return goals 540 | 541 | end Tactic 542 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Mutual induction tactic for Lean 4 2 | 3 | * [Mutual induction with recursors](#mutual-induction-with-recursors) 4 | * [Mutual induction with the tactic](#mutual-induction-with-the-tactic) 5 | * [How does the tactic work?](#how-does-the-tactic-work) 6 | 1. [Compute targets and generalized variables](#1-compute-targets-and-generalized-variables) 7 | 2. [Check coverage and variable scoping](#2-check-coverage-and-variable-scoping) 8 | 3. [Generalize variables and compute motives](#3-generalize-variables-and-compute-motives) 9 | 4. [Apply recursors](#4-apply-recursors) 10 | 5. [Deduplicate subgoals](#5-deduplicate-subgoals) 11 | 6. [Add subgoals](#6-add-subgoals) 12 | * [How do joint theorems work?](#how-do-joint-theorems-work) 13 | * [Extensions](#extensions) 14 | * [Transposed clauses](#transposed-clauses) 15 | * [`with` clauses](#with-clauses) 16 | * [Mutual induction in other proof assistants](#mutual-induction-in-other-proof-assistants) 17 | * [Rocq](#rocq) 18 | * [Isabelle](#isabelle) 19 | 20 | --- 21 | 22 | This is an experimental mutual induction tactic that acts on multiple goals. 23 | For now, the syntax looks like the below, 24 | with no support yet for the usual `induction` tactic's features `with` 25 | or for generalizing different variables in each goal. 26 | 27 | ```lean 28 | mutual_induction x₁, ..., xₙ (using r₁, ..., rₙ)? (generalizing y₁ ... yₘ)? 29 | ``` 30 | 31 | Generating multiple goals can be done using existing declarations refinement tactics, 32 | but a convenient `joint theorem` declaration form is also provided. 33 | 34 | ```lean 35 | joint.{u,...}? ((y : A)...)? 36 | theorem thm₁ x₁... : B₁ 37 | ... 38 | theorem thmₙ xₙ... : Bₙ 39 | by [tactics] 40 | ``` 41 | 42 | The doc comment for the tactic gives an example using mutual even/odd naturals, 43 | which demonstrates basic usage but not some of the subtler issues with implementation. 44 | Here, we'll use even/odd predicates over naturals as the running example. 45 | 46 | ```lean 47 | mutual 48 | inductive Even : Nat → Prop 49 | | zero : Even 0 50 | | succ : ∀ n, Odd n → Even (n + 1) 51 | inductive Odd : Nat → Prop 52 | | succ : ∀ n, Even n → Odd (n + 1) 53 | end 54 | open Even Odd 55 | ``` 56 | 57 | ## Mutual induction with recursors 58 | 59 | Recursors are generated for this mutual pair of inductives, 60 | which share the same motives and cases but with different conclusions. 61 | 62 | ```lean 63 | Even.rec : ∀ {motive_1 : ∀ n, Even n → Prop} {motive_2 : ∀ n, Odd n → Prop}, 64 | -- cases for Even 65 | motive_1 0 zero → 66 | (∀ n (on : Odd n), motive_2 n on → motive_1 (n + 1) (succ n on)) → 67 | -- cases for Odd 68 | (∀ n (en : Even n), motive_1 n en → motive_2 (n + 1) (succ n en)) → 69 | -- conclusion for Even 70 | ∀ {n} (en : Even n), motive_1 n en 71 | 72 | Odd.rec : ∀ {motive_1 : ∀ n, Even n → Prop} {motive_2 : ∀ n, Odd n → Prop}, 73 | -- cases for Even 74 | motive_1 0 zero → 75 | (∀ n (on : Odd n), motive_2 n on → motive_1 (n + 1) (succ n on)) → 76 | -- cases for Odd 77 | (∀ n (en : Even n), motive_1 n en → motive_2 (n + 1) (succ n en)) → 78 | -- conclusion for Odd 79 | ∀ {n} (on : Odd n), motive_2 n on 80 | ``` 81 | 82 | We can conjoin the two conclusions into a single recursor using these two 83 | to avoid duplicating work proving the identical cases. 84 | 85 | ```lean 86 | theorem rec : ∀ {motive_1 motive_2}, ... → 87 | (∀ {n} (en : Even n), motive_1 n en) ∧ 88 | (∀ {n} (on : Odd n), motive_2 n on) := by ... 89 | ``` 90 | 91 | The disadvantage of using the conjoined recursor 92 | is that the goal must have exactly this shape 93 | for unification to solve goals automatically. 94 | Otherwise, it must be manually applied and manipulated. 95 | For example, reordering and pulling the `n` out of the conjunction is equivalent: 96 | 97 | ```lean 98 | theorem rec' : ∀ {motive_1 motive_2}, ... → ∀ {n}, 99 | (∀ (on : Odd n), motive_2 n on) ∧ 100 | (∀ (en : Even n), motive_1 n en) := by ... 101 | ``` 102 | 103 | but one cannot be proven from the other without destructing a conjunction 104 | along with additional instantiations or introductions. 105 | 106 | ## Mutual induction with the tactic 107 | 108 | The aim of this mutual induction tactic is to alleviate this manipulation tedium 109 | by avoiding conjunctions altogether. 110 | To demonstrate, we prove an inversion theorem about parity of addition: 111 | if the addition of two naturals is even, then they are either both even or both odd; 112 | and if the addition of two naturals is odd, then one must be even and the other odd. 113 | 114 | ```lean 115 | joint (n : Nat) (m : Nat) 116 | theorem evenAddInv (onm : Even (n + m)) : (Even n ∧ Even m) ∨ (Odd n ∧ Odd m) 117 | theorem oddAddInv (enm : Odd (n + m)) : (Odd n ∧ Even m) ∨ (Even n ∧ Odd m) 118 | by 119 | case' evenAddInv => generalize e₂ : n + m = k₂ at onm 120 | case' oddAddInv => generalize e₁ : n + m = k₁ at enm 121 | ``` 122 | 123 | We wish induct on the proofs of `Even (n + m)` and `Odd (n + m)`. 124 | Just as with the usual `induction` tactic, 125 | we can't induct on inductives whose indices aren't variables, 126 | so we generalize `n + m` over an equality. 127 | The proof state now looks like the below. 128 | 129 | ```lean 130 | ▼ case evenAddInv 131 | n m k₁ : Nat 132 | e₁ : n + m = k₁ 133 | enm : Even k₁ 134 | ⊢ (Even n ∧ Even m) ∨ (Odd n ∧ Odd m) 135 | 136 | ▼ case oddAddInv 137 | n m k₂ : Nat 138 | e₂ : n + m = k₂ 139 | onm : Odd k₂ 140 | ⊢ (Odd n ∧ Even m) ∨ (Even n ∧ Odd m) 141 | ``` 142 | 143 | We now apply mutual induction by `mutual_induction enm, onm generalizing n`, 144 | which says that we are doing induction on `enm` in goal `left` and on `onm` in goal `right`. 145 | It yields the following goals. 146 | 147 | ```lean 148 | ▼ case evenAddInv.zero 149 | m n : Nat 150 | e₁ : n + m = 0 151 | ⊢ (Even n ∧ Even m) ∨ (Odd n ∧ Odd m) 152 | 153 | ▼ case evenAddInv.succ 154 | m k₁ : Nat 155 | ok : Odd k₁ 156 | ih : ∀ (n : Nat), n + m = k₁ → (Odd n ∧ Even m) ∨ (Even n ∧ Odd m) 157 | n : Nat 158 | e₁ : n + m = k₁ + 1 159 | ⊢ (Even n ∧ Even m) ∨ (Odd n ∧ Odd m) 160 | 161 | ▼ case oddAddInv.succ 162 | m k₂ : Nat 163 | ek : Even k₂ 164 | ih : ∀ (n : Nat), n + m = k₂ → (Even n ∧ Even m) ∨ (Odd n ∧ Odd m) 165 | n : Nat 166 | e₂ : n + m = k₂ + 1 167 | ⊢ (Odd n ∧ Even m) ∨ (Even n ∧ Odd m) 168 | ``` 169 | 170 | The proof proceeds by cases on `n` in the `succ` cases, 171 | which is why we generalize the induction hypothesis over it. 172 | The full proof can be found at `EvenOdd.plusEvenOdd`. 173 | 174 | ## How does the tactic work? 175 | 176 | The tactic proceeds in stages: 177 | 178 | 1. Compute whatever information we can from each goal independently. 179 | 2. Ensure that the goals match the mutual inductives 180 | and share the generalized variables. 181 | 3. Compute more information from all goals in tandem. 182 | 4. Apply the appropriate recursor for each goal. 183 | 5. Combine duplicate subgoals from the recursors. 184 | 6. Add subgoals to the proof state, (re)introducing variables. 185 | 186 | ### 1. Compute targets and generalized variables 187 | 188 | The user specifies that the target of `evenAddInv` is `enm` and the target of `oddAddInv` is `onm`. 189 | However, the indices of their types are considered as auxiliary targets 190 | because the motives need to abstract over them as well. 191 | This would be `k₁` and `k₂` in their respective goals. 192 | 193 | We also retrieve the inductive type to which the primary target belongs, 194 | along with the positions of the motive for that inductive type and of the targets 195 | within the type of the provided recursor. 196 | If no custom recursor is provided, the default recursor for that inductive type is used. 197 | In our example, for the default recursors, 198 | these are `Even`, `[0, 5, 6]` for the first goal, 199 | and `Odd`, `[1, 5, 6]` for the second. 200 | 201 | Now, we need to find the variables to generalize the goals over, which are 202 | 203 | 1. The variables whose types depend on the targets; but also 204 | 2. In the types of *those* variables, the other variables they depend on; and 205 | 3. The other variables that the goal depends on. 206 | 207 | The variables in bucket (1) are what get generalized in the usual `induction` tactic. 208 | However, because we are dealing with multiple goals in different contexts, 209 | we need to ensure when we turn the goals into motives that they are closed, 210 | which means possibly generalizing over all variables in buckets (2) and (3). 211 | In our example, this corresponds to 212 | 213 | * (1) `e₁` and `e₂` by dependency on auxiliary targets `k₁` and `k₂`, respectively; 214 | * (2, 3) `n` and `m` depended upon by `e₁`, `e₂`, and the goals. 215 | 216 | This work is done by `Lean.Elab.Tactic.getSubgoal`. 217 | 218 | ### 2. Check coverage and variable scoping 219 | 220 | Although the previous step ensures that the targets are inductive, 221 | we also need to ensure that 222 | 223 | * The targets all belong to inductive types in the *same* mutual declaration; 224 | * The targets each belong to a *different* inductive type; and 225 | * ~~The targets belong to *all* of the mutual inductive types.~~ 226 | 227 | These conditions ensure that we have the correct motives needed 228 | to apply the recursors to each goal. 229 | If a motive is missing due to a missing target for one of the inductive types, 230 | then ~~we add that motive as an additional goal~~ 231 | it gets set to the trivial constant type `PUnit`. 232 | 233 | We then check that the provided generalized variables 234 | are indeed shared across goals, i.e. declared in each of the goals' contexts. 235 | Variables depended upon that aren't shared across goals must always be generalized 236 | to produce closed motives, as described in the previous step. 237 | 238 | This work is done by `Lean.Elab.Tactic.checkTargets` 239 | and by `Lean.Elab.Tactic.checkFVars`. 240 | 241 | ### 3. Generalize variables and compute motives 242 | 243 | Although we compute the variables that may be generalized independently for each goal, 244 | we don't yet actually generalize them, 245 | because there may variables that happen to be common to all goals 246 | that don't need to be generalized over because they'll always be in scope. 247 | In this case, because `n` is explicitly generalized, `m` is the only such variable, 248 | which makes sense because it was introduced outside of the conjunction. 249 | Only now do we finally generalize the variables 250 | and compute the motives by abstracting the goals over the targets. 251 | 252 | ```lean 253 | ▼ case evenAddInv 254 | m k₁ : Nat 255 | enm : Even k₁ 256 | motive_1 := λ (k₁ : Nat) (enm : Even k₁) ↦ 257 | ∀ (n : Nat) (e₁ : n + m = 0), (Even n ∧ Even m₁) ∨ (Odd n ∧ Odd m₁) 258 | ⊢ ∀ (n : Nat) (e₁ : n + m = 0), (Even n ∧ Even m₁) ∨ (Odd n ∧ Odd m₁) 259 | 260 | ▼ case oddAddInv 261 | m k₂ : Nat 262 | onm : Odd k₂ 263 | motive_2 := λ (k₂ : Nat) (onm : Odd k₂) ↦ 264 | ∀ (n : Nat) (e₂ : n + m = 0), (Odd n ∧ Even m₂) ∨ (Even n ∧ Odd m₂) 265 | ⊢ ∀ (n : Nat) (e₂ : n + m = 0), (Odd n ∧ Even m₂) ∨ (Even n ∧ Odd m₂) 266 | ``` 267 | 268 | For each goal, we know the position of the motive that applies to its target 269 | within the type of the recursor to apply, 270 | but not the positions of the other motives. 271 | Therefore, we assume that all recursors take all motives in exactly the same order, 272 | which is true of the autogenerated recursors, 273 | and sort the motives in that order. 274 | This means that custom recursors provided via `using` 275 | must all quantify over motives in exactly the same order. 276 | Otherwise, finding the correct positions of all motives in all goals 277 | would require O(n²) type comparisons. 278 | 279 | This work is done by `Lean.Elab.Tactics.addMotives`. 280 | 281 | ### 4. Apply recursors 282 | 283 | In each goal, we know the primary target and its inductive type, 284 | so we can retrieve the recursor and instantiate it with the motives and targets, 285 | leaving the remaining arguments as subgoals to be solved. 286 | 287 | ```lean 288 | ▼ case evenAddInv 289 | m k₁ : Nat 290 | enm : Even k₁ 291 | motive_1 := λ (k₁ : Nat) (enm : Even k₁) ↦ 292 | ∀ (n : Nat) (e₁ : n + m = 0), (Even n ∧ Even m₁) ∨ (Odd n ∧ Odd m₁) 293 | motive_2 := λ (k₂ : Nat) (onm : Odd k₂) ↦ 294 | ∀ (n : Nat) (e₂ : n + m = 0), (Odd n ∧ Even m₂) ∨ (Even n ∧ Odd m₂) 295 | ⊢ @Even.rec motive_1 motive_2 ?evenAddInv.Even.zero ?evenAddInv.Even.succ ?evenAddInv.Odd.succ k₁ enm 296 | : ∀ (n : Nat) (e₁ : n + m = 0), (Even n ∧ Even m) ∨ (Odd n ∧ Odd m) 297 | 298 | ▼ case oddAddInv 299 | m k₂ : Nat 300 | onm : Odd k₂ 301 | motive_1 := λ (k₁ : Nat) (enm : Even k₁) ↦ 302 | ∀ (n : Nat) (e₁ : n + m = 0), (Even n ∧ Even m₁) ∨ (Odd n ∧ Odd m₁) 303 | motive_2 := λ (k₂ : Nat) (onm : Odd k₂) ↦ 304 | ∀ (n : Nat) (e₂ : n + m = 0), (Odd n ∧ Even m₂) ∨ (Even n ∧ Odd m₂) 305 | ⊢ @Odd.rec motive_1 motive_2 ?oddAddInv.Even.zero ?oddAddInv.Even.succ ?oddAddInv.Odd.succ k₁ enm 306 | : ∀ (n : Nat) (e₂ : n + m = 0), (Odd n ∧ Even m) ∨ (Even n ∧ Odd m) 307 | ``` 308 | 309 | The subgoals are collected up as a 2D array. 310 | 311 | ```lean4 312 | [[?evenAddInv.Even.zero, ?evenAddInv.Even.succ, ?evenAddInv.Odd.succ], 313 | [?oddAddInv.Even.zero, ?oddAddInv.Even.succ, ?oddAddInv.Odd.succ]] 314 | ``` 315 | 316 | This work is done by `Lean.Elab.Tactic.evalSubgoal`. 317 | 318 | ### 5. Deduplicate subgoals 319 | 320 | By virtue of the types of the recursors, 321 | the arrays of subgoals have the same type pointwise, 322 | so we can pick one from each array and equate the rest to it. 323 | These subgoals each correspond to one of the constructors of the mutual inductive types. 324 | We intuitively expect the prefix of the name of the subgoal for a particular constructor 325 | to match the original goal whose target's inductive type contains that constructor. 326 | Picking the subgoals that prove the motive that applies to the parent goal's target 327 | ensures that we get the correct name. 328 | Once they have been picked, the inductive type's name can be removed from the subgoal's name. 329 | 330 | ```lean 331 | ▼ case evenAddInv.Even.zero 332 | m : Nat 333 | ⊢ ∀ (n : Nat) (e₁ : n + m = 0), (Even n ∧ Even m) ∨ (Odd n ∧ Odd m) 334 | 335 | ▼ case evenAddInv.Even.succ 336 | m : Nat 337 | ⊢ ∀ (k₁ : Nat) (enm : Even k₁), 338 | (∀ (n : Nat), n + m = k₁ → (Odd n ∧ Even m) ∨ (Even n ∧ Odd m)) → 339 | ∀ (n : Nat) (e₁ : n + m = k₁ + 1), (Even n ∧ Even m) ∨ (Odd n ∧ Odd m) 340 | 341 | ▼ case oddAddInv.Odd.succ 342 | m : Nat 343 | ⊢ ∀ (k₁ : Nat) (enm : Even k₁), 344 | (∀ (n : Nat), n + m = k₂ → Even n ∧ Even m ∨ Odd n ∧ Odd m) → 345 | ∀ (n : Nat) (e₂ : n + m = k₂ + 1), (Odd n ∧ Even m) ∨ (Even n ∧ Odd m) 346 | 347 | ▶ case oddAddInv.Even.zero := evenAddInv.Even.zero 348 | ▶ case oddAddInv.Even.succ := evenAddInv.Even.succ 349 | ▶ case evenAddInv.Odd.succ := oddAddInv.Odd.succ 350 | ``` 351 | 352 | This work is done by `Lean.Elab.Tactic.deduplicate`. 353 | 354 | ### 6. Add subgoals 355 | 356 | Each subgoal is added to the proof state's list of goals, 357 | but only after (re)introducing into the context 358 | the constructor arguments, the induction hypotheses, and the generalized variables. 359 | If the goal is trivially `PUnit`, it gets solved by its constructor. 360 | If an induction hypothesis is the trivial `PUnit`, it gets cleared away. 361 | 362 | ```lean 363 | ▼ case evenAddInv.zero 364 | m n : Nat 365 | e₁ : n + m = 0 366 | ⊢ (Even n ∧ Even m) ∨ (Odd n ∧ Odd m) 367 | 368 | ▼ case evenAddInv.succ 369 | m k₁ : Nat 370 | enm : Even k₁ 371 | ih : ∀ (n : Nat), n + m = k₁ → (Odd n ∧ Even m) ∨ (Even n ∧ Odd m) 372 | n : Nat 373 | e₁ : n + m = k₁ + 1 374 | ⊢ (Even n ∧ Even m) ∨ (Odd n ∧ Odd m) 375 | 376 | ▼ case oddAddInv.succ 377 | m k₁ : Nat 378 | enm : Even k₁ 379 | ih : ∀ (n : Nat), n + m = k₂ → Even n ∧ Even m ∨ Odd n ∧ Odd m 380 | n : Nat 381 | e₂ : n + m = k₂ + 1 382 | ⊢ (Odd n ∧ Even m) ∨ (Even n ∧ Odd m) 383 | ``` 384 | 385 | This work is done by `Lean.Elab.Tactic.addSubgoal`. 386 | 387 | ## How do joint theorems work? 388 | 389 | The general form of the joint theorem command is as follows: 390 | 391 | ```lean 392 | joint.{u,...} (y : A)... 393 | theorem thm₀ (x₀ : B₀)... : C₀ 394 | ... 395 | theorem thmₙ (xₙ : Bₙ)... : Cₙ 396 | by [tactics] 397 | ``` 398 | 399 | The optional universe parameters `u, ...` and variables `y, ...` 400 | are bound within all of the theorems' types declarations. 401 | At the moment, shared universe parameters *must* be used in all theorems, 402 | as they will otherwise throw unused universe parameters errors. 403 | As usual, the binders `x₀..., ..., xₙ...` may also be (strict) implicit `{x : B}`, `⦃x : B⦄`, 404 | and the type annotation `B` may be omitted. 405 | The variables `y, ...` may similarly be (strict) implicit, 406 | but the type annotation `A` may *not* be omitted, 407 | as the way the command is expanded makes inferring their types difficult. 408 | 409 | The command first expands into an array of theorems and their proofs in proof mode. 410 | This definition's name is generated from the theorems' names, 411 | and is currently `_thms_ ≡ _thm₀_..._thmₙ_`. 412 | 413 | ```lean 414 | def _thms_.{u,...} (y : A)... : Array (Σ P : Sort _, P) := by 415 | refine #[ 416 | ⟨(∀ (x₁ : B₁)..., C₁), (λ (x₁ : B₁)... ↦ ?thm₀)⟩, 417 | ... 418 | ⟨(∀ (xₙ : Bₙ)..., Cₙ), (λ (xₙ : Bₙ)... ↦ ?thmₙ)⟩ 419 | ] 420 | [tactics] 421 | ``` 422 | 423 | At the tactics' position, there is one goal named after each theorem, 424 | whose context contains the shared and bound variables, 425 | and whose type is the theorem's type. 426 | 427 | Subsequent theorem definitions are created for each theorem declaration, 428 | indexing into the array for their proofs. 429 | 430 | ```lean 431 | def thm₀.{u,...} (y : A)... : ∀ (x₀ : B₀)..., C₀ := (@_thms_.{u,...} y...)[0].snd 432 | ... 433 | def thmₙ.{u,...} (y : A)... : ∀ (xₙ : Bₙ)..., Cₙ := (@_thms_.{u,...} y...)[n].snd 434 | ``` 435 | 436 | ## Extensions 437 | 438 | There are many possible usability improvements for this tactic. 439 | A simple one is adding an option to keep missing motives as additional goals 440 | instead of filling them with `PUnit`, but this use case seems rare. 441 | Below are other more elaborate potential extensions. 442 | 443 | ### Transposed clauses 444 | 445 | The tactic's current syntax doesn't allow for 446 | generalizing different sets of variables in each goal, 447 | only generalizing the same set in all goals. 448 | We can imagine using comma-separated sets of generalized variables, 449 | just as the `using` clause takes comma-separated recursors. 450 | 451 | ```lean 452 | mutual_induction x₁, ..., xₙ 453 | (using r₁, ..., rₙ)? 454 | (generalizing z₁..., ..., zₙ...)? 455 | ``` 456 | 457 | While `using` can *always* take recursors, 458 | `generalizing` might have goals with *no* generalized variables. 459 | Furthermore, generalizing the same variable in all goals 460 | now requires repeating that variable in all subclauses. 461 | This leads to slightly awkward notation in some situations. 462 | 463 | ```lean 464 | mutual_induction x₁, x₂, x₃, x₄, x₅ generalizing ,,,,z 465 | mutual_induction x₁, x₂, x₃, x₄, x₅ generalizing z, z, z, z, z 466 | ``` 467 | 468 | An alternate syntax would be to transpose the goals, 469 | specifying the clauses for each one separately. 470 | We can refer to the goals by name this way, 471 | which frees us from specifying the targets in the fixed given order. 472 | 473 | ```lean 474 | mutual_induction 475 | | goal₁ => x₁ (using r₁)? (generalizing z₁...)? 476 | ... 477 | | goalₙ => xₙ (using rₙ)? (generalizing zₙ...)? 478 | ``` 479 | 480 | This is quite a bit wordier, 481 | but this could be added as an alternate expanded syntax 482 | alongside the current terser one. 483 | 484 | ### `with` clauses 485 | 486 | At the moment, the tactic is missing support for `with` clauses. 487 | The `with` clause applies tactics to the specified subgoals generated, 488 | and is sugar for subsequent `case` expressions. 489 | The syntax for the tactic with this clause might look something like the following. 490 | 491 | ```lean 492 | mutual_induction x₁, ..., xₙ with 493 | | tag₁ z₁ ... zᵢ₁ => tac₁ 494 | ... 495 | | tagₖ z₁ ... zᵢₖ => tacₖ 496 | ``` 497 | 498 | Combined with the transposed syntax from above, though, might be a bit unwieldy. 499 | 500 | ```lean 501 | mutual_induction 502 | | goal₁ => x₁ generalizing z₂ z₃ 503 | | goal₂ => x₂ generalizing z₁ z₃ 504 | | goal₃ => x₃ generalizing z₁ z₂ 505 | with 506 | | goal₁.tag₁ y => tac₁ 507 | | goal₁.tag₂ y => tac₂ 508 | | goal₂.tag₁ y => tac₃ 509 | | goal₃.tag₁ y => tac₄ 510 | ``` 511 | 512 | ## Mutual induction in other proof assistants 513 | 514 | ### Rocq 515 | 516 | By default, Rocq generates nonmutual induction principles for mutual inductives. 517 | For instance, the even/odd predicates 518 | 519 | ```rocq 520 | Inductive Even : nat -> Prop := 521 | | zero : Even 0 522 | | succ_odd n : Odd n -> Even (S n) 523 | with Odd : nat -> Prop := 524 | | succ_even n : Even n -> Odd (S n). 525 | ``` 526 | 527 | come with induction principles defined via `fix` over independent motives. 528 | 529 | ```rocq 530 | Even_ind : forall P : nat -> Prop, 531 | P 0 -> 532 | (forall n, Odd n -> P (S n)) -> 533 | forall n, Even n -> P n 534 | 535 | Odd_ind : forall P : nat -> Prop, 536 | (forall n, Even n -> P (S n)) -> 537 | forall n, Odd -> P n 538 | ``` 539 | 540 | Using `Scheme` generates the mutual induction principles we expect. 541 | 542 | ```rocq 543 | Scheme Even_Odd_ind := Induction for Even Sort Prop 544 | with Odd_Even_ind := Induction for Odd Sort Prop. 545 | ``` 546 | 547 | ```rocq 548 | Even_Odd_ind : forall (P : forall n, Even n -> Prop) (P0 : forall n, Odd n -> Prop) 549 | P 0 -> 550 | (forall n (o : Odd n), P0 n 0 -> P (S n) (succ_odd n o)) -> 551 | (forall n (e : Odd n), P n e -> P0 (S n) (odd_succ n e)) -> 552 | forall n (e : Even n), P n e 553 | 554 | Odd_Even_ind : forall (P : forall n, Even n -> Prop) (P0 : forall n, Odd n -> Prop) 555 | P 0 -> 556 | (forall n (o : Odd n), P0 n 0 -> P (S n) (succ_odd n o)) -> 557 | (forall n (e : Odd n), P n e -> P0 (S n) (odd_succ n e)) -> 558 | forall n (o : Odd n), P0 n o 559 | ``` 560 | 561 | These can be used with the induction tactic just as in Lean: 562 | `induction e using Even_Odd_ind with (P0 := Q)` 563 | in Rocq corresponds roughly to 564 | `induction e using Even.rec (motive_2 := Q)` 565 | in Lean. 566 | 567 | But we can do one better: `Combined Scheme` conjoins the conclusions 568 | and deduplicates the cases, similar to what `mutual_induction` does internally. 569 | 570 | ```rocq 571 | Combined Scheme Even_Odd_mutind from Even_Odd_ind, Odd_Even_ind. 572 | ``` 573 | 574 | ```rocq 575 | Even_Odd_mutind : forall (P : forall n, Even n -> Prop) (P0 : forall n, Odd n -> Prop) 576 | P 0 -> 577 | (forall n (o : Odd n), P0 n 0 -> P (S n) (succ_odd n o)) -> 578 | (forall n (e : Odd n), P n e -> P0 (S n) (odd_succ n e)) -> 579 | (forall n (e : Even n), P n e) /\ (forall n (o : Odd n), P0 n o) 580 | ``` 581 | 582 | This combined scheme, although conveniently generated for us, 583 | has the same disadvantages as outlined in the beginning. 584 | 585 | Additionally, Rocq provides top-level mutual theorem statements, 586 | which generates two goals where both theorems are in scope. 587 | 588 | ```rocq 589 | Lemma plusEven {n m} (e : Even (n + m)) : (Even n /\ Even m) \/ (Odd n /\ Odd m) 590 | with plusOdd {n m} (o : Odd (n + m)) : (Odd n /\ Even m) \/ (Even n /\ Odd m)). 591 | ``` 592 | 593 | ```rocq 594 | plusEven : forall n m, Even (n + m) -> (Even n /\ Even m) \/ (Odd n /\ Odd m) 595 | plusOdd : forall n m, Odd (n + m) -> (Odd n /\ Even m) \/ (Even n /\ Odd m)) 596 | n, m : nat 597 | ------------------------------------------- (1/2) 598 | (Even n /\ Even m) \/ (Odd n /\ Odd m) 599 | ------------------------------------------- (2/2) 600 | (Odd n /\ Even m) \/ (Even n /\ Odd m)) 601 | ``` 602 | 603 | The idea with this setup is to be able to prove these lemmas by mutual recursion, 604 | explicitly calling the other lemma on mutual subarguments. 605 | The danger here is the same as proofs by mutual recursion in Lean: 606 | the user has to make sure to only instantiate mutual lemmas on structurally smaller arguments, 607 | or else specify a different termination metric altogether. 608 | 609 | ### Isabelle 610 | 611 | *I am not an Isabelle user and haven't tested any code, 612 | so this information may be inaccurate.* 613 | 614 | The Isar proof language for Isabelle provides an induction tactic 615 | that appears to work adequately on mutual inductives as well. 616 | Considering again the even/odd predicates, 617 | the inversion lemma on parity of addition is stated mutually. 618 | 619 | ```isabelle 620 | inductive even and odd where 621 | "even 0" 622 | | "even n ⟹ odd (n + 1)" 623 | | "odd n ⟹ even (n + 1)" 624 | 625 | lemma 626 | fixes n m :: 'a 627 | shows 628 | "even (n + m) ⟹ (even n ∧ even m) ∨ (odd n ∧ odd m)" and 629 | "odd (n + m) ⟹ (odd n ∧ even m) ∨ (even m ∧ odd m)" 630 | ``` 631 | 632 | The proof proceeds by applying induction, which has the general form 633 | (assuming two goals, but more can be conjoined by `and`): 634 | 635 | ```isabelle 636 | induct (x = t)... and (y = u)... arbitrary: v... and w... rule: R 637 | ``` 638 | 639 | corresponding roughly to first `generalize t = x; ...` and `generalize u = y; ...` 640 | on the antecedents of the first and second goals respectively, 641 | then induction on those antecedents generalizing `v...` and `w...` respectively, 642 | using the induction principle `R`. 643 | Even if generalization is not needed, 644 | the indices of each inductive at the very least need to all be provided. 645 | So to continue our proof above, 646 | 647 | ```isabelle 648 | apply ( 649 | induct (k = n + m) and (k = n + m) 650 | arbitrary: n and n 651 | rule: even_odd.inducts 652 | ) 653 | ``` 654 | 655 | should advance the proof to the three cases for each constructor. 656 | The advantage of this `and` syntax, in particular for `arbitrary`, 657 | is that different goals can have different arguments generalized; 658 | the disadvantage is that if all goals generalize the same argument, 659 | they still must be specified for each goal. 660 | Additionally, if only the last of four goals needs an argument generalized, 661 | it would require writing `arbitrary: and and and n`. 662 | 663 | Reference: [The Isabelle/Isar Reference Manual](https://isabelle.in.tum.de/doc/isar-ref.pdf#subsection.6.5.2) 664 | 665 | # Towards nested induction: `mk_all` 666 | 667 | As an extra bonus, this repo also includes a `mk_all` attribute for inductive types 668 | that generates new definitions that lift predicates over the given parameters 669 | to predicates over that inductive type. 670 | For example, consider the following list type that we want to lift predicates over. 671 | 672 | ```lean 673 | @[mk_all α] 674 | inductive List (α : Type) where 675 | | nil : List α 676 | | cons : α → List α → List α 677 | ``` 678 | 679 | By default, `List` is in a `Type` universe. 680 | Then two definitions are generated: `List.all` and `List.iall`, 681 | which lift predicates over `α` to predicates over `List α` 682 | into `Prop` and into `Type`, respectively. 683 | Internally, the predicates are implemented using the list recursor, 684 | but are equivalent to the following recursive functions. 685 | 686 | ```lean 687 | def List.all {α : Type} (P : α → Type) : List α → Type 688 | | nil => Unit 689 | | cons x xs => P x × List.all P xs 690 | 691 | def List.iall {α : Type} (P : α → Prop) : List α → Prop 692 | | nil => True 693 | | cons x xs => P x ∧ List.iall P xs 694 | ``` 695 | 696 | If `List` is in the `Prop` universe, 697 | then only a `List.all` into `Prop` is generated as an inductive type. 698 | 699 | ```lean 700 | inductive List.all {α : Type} (P : α → Prop) : Lst α → Prop where 701 | | nil : all P nil 702 | | cons x xs : P x → all P xs → all P (cons x xs) 703 | ``` 704 | 705 | Given some list `xs = [a, b, c, ...]`, 706 | the predicate `List.all P xs` represents a list of propositions 707 | `[P a, P b, P c, ...]`. 708 | This is particularly useful if we have an inductive nested inside of lists, 709 | and we want to lift a predicate over the nested inductive to a predicate on lists of them. 710 | The classic example is the rose tree. 711 | 712 | ```lean 713 | inductive Tree (α : Type) where 714 | | leaf : Tree α 715 | | node : α → List (Tree α) → Tree α 716 | ``` 717 | 718 | The recursor for `Tree` treats the nesting opaquely, 719 | demanding a predicate over `List (Tree α)` as a second motive, 720 | and generating additional cases for the nesting inductive `List`. 721 | 722 | ```lean 723 | Tree.rec : ∀ {α : Type} {motive_1 : Tree α → Type} {motive_2 : List (Tree α) → Type}, 724 | -- cases for Tree 725 | motive_1 leaf → 726 | (∀ x xs, motive_2 xs → motive_1 (node x xs)) → 727 | -- cases for List 728 | motive_2 nil → 729 | (∀ x xs, motive_1 x → motive_2 xs → motive_2 (cons x xs)) → 730 | -- conclusion for Tree 731 | ∀ (t : Tree α), motive_1 t 732 | ``` 733 | 734 | What would be more useful is if `motive_1` were applied to each subtree in the node, 735 | which is exactly what `List.all motive_1` does. 736 | Such elimination principles into `Prop` and `Type` can easily be proven. 737 | 738 | ```lean 739 | theorem Tree.ielim {α} (P : Tree α → Prop) 740 | (hleaf : P leaf) 741 | (hnode : ∀ {x xs}, List.iall P xs → P (node x xs)) : 742 | ∀ t, P t := by 743 | intro t 744 | induction t using Tree.rec (motive_2 := List.iall P) 745 | case leaf => exact hleaf 746 | case node ih => exact hnode ih 747 | case nil => exact ⟨⟩ 748 | case cons hx hxs => exact ⟨hx, hxs⟩ 749 | 750 | noncomputable def Tree.elim {α} (P : Tree α → Type) 751 | (hleaf : P leaf) 752 | (hnode : ∀ {x xs}, List.all P xs → P (node x xs)) : 753 | ∀ t, P t := by 754 | intro t 755 | induction t using Tree.rec (motive_2 := List.all P) 756 | case leaf => exact hleaf 757 | case node ih => exact hnode ih 758 | case nil => exact ⟨⟩ 759 | case cons hx hxs => exact ⟨hx, hxs⟩ 760 | ``` 761 | 762 | Unfortunately, the recursors for nested inductives in Lean are currently noncomputable, 763 | so we can at most define a better noncomputable elimination principle for `Tree`. 764 | 765 | In general, the `mk_all` attribute can be applied with multiple parameters, 766 | so long as each parameter appears strictly positively in the inductive definition. 767 | Lean doesn't support nonuniform parameters, 768 | so generating predicate liftings for deeply nested inductives 769 | (such as perfect trees) is not supported either. 770 | Future work may include the following: 771 | 772 | * For inductives not in `Prop`, 773 | making `simp` reduce `.(i)all` on constructors 774 | as if they were defined as the equivalent recursive definition. 775 | * Generating `.(i)all` for already-defined inductive types from other modules. 776 | * Generating the definitions of the `.(i)all`-augmented eliminators automatically. 777 | * Generating corresponding `.below` and `.brecOn` definitions 778 | that incorporate `.(i)all`. 779 | 780 | The design of `Lean.Meta.MkAll.mkAllDef` and `Lean.Meta.MkAll.mkIAllDef` 781 | are based on `Lean.Meta.BRecOn.mkBelowOrIBelow` and `Lean.Meta.IndPredBelow.mkBelow` respectively, 782 | so the proof search for the eliminators will likely follow the design of 783 | `Lean.Meta.BRecOn.mkBRecOn` and `Lean.Meta.BRecOn.mkBInductionOn`. 784 | --------------------------------------------------------------------------------