├── .github ├── cnf_tests.txt ├── fof_tests.txt ├── tff_tests.txt ├── thf_tests.txt └── workflows │ ├── test.yml │ └── test_exe.yml ├── .gitignore ├── Duper.lean ├── Duper ├── BackwardSimplification.lean ├── Clause.lean ├── ClauseStreamHeap.lean ├── DUnif │ ├── Bindings.lean │ ├── DApply.lean │ ├── DRefl.lean │ ├── Oracles.lean │ ├── Readme.md │ ├── UnifProblem.lean │ ├── UnifRules.lean │ └── Utils.lean ├── DiscrTree.lean ├── Expr.lean ├── Fingerprint.lean ├── Interface.lean ├── MClause.lean ├── Match.lean ├── Order.lean ├── Preprocessing.lean ├── ProofReconstruction.lean ├── ProverM.lean ├── RuleM.lean ├── Rules │ ├── ArgumentCongruence.lean │ ├── BetaEtaReduction.lean │ ├── BoolHoist.lean │ ├── BoolSimp.lean │ ├── ClauseSubsumption.lean │ ├── Clausification.lean │ ├── ClausifyPropEq.lean │ ├── ContextualLiteralCutting.lean │ ├── DatatypeAcyclicity.lean │ ├── DatatypeDistinctness.lean │ ├── DatatypeExhaustiveness.lean │ ├── DatatypeInjectivity.lean │ ├── DecElim.lean │ ├── Demodulation.lean │ ├── DestructiveEqualityResolution.lean │ ├── ElimDupLit.lean │ ├── ElimResolvedLit.lean │ ├── EqHoist.lean │ ├── EqualityFactoring.lean │ ├── EqualityResolution.lean │ ├── EqualitySubsumption.lean │ ├── ExistsHoist.lean │ ├── FalseElim.lean │ ├── FluidBoolHoist.lean │ ├── FluidSup.lean │ ├── ForallHoist.lean │ ├── IdentBoolFalseElim.lean │ ├── IdentBoolHoist.lean │ ├── IdentPropFalseElim.lean │ ├── NeHoist.lean │ ├── SimplifyReflect.lean │ ├── Superposition.lean │ ├── SyntacticTautologyDeletion1.lean │ ├── SyntacticTautologyDeletion2.lean │ └── SyntacticTautologyDeletion3.lean ├── Saturate.lean ├── Selection.lean ├── Simp.lean ├── SubsumptionTrie.lean ├── TPTP.lean ├── TPTPParser │ ├── MacroDecl.lean │ ├── PrattParser.lean │ └── SyntaxDecl.lean ├── Tactic.lean ├── Tests │ ├── DApplyTests.lean │ ├── DReflTests.lean │ ├── bugs.lean │ ├── casc29_tests.lean │ ├── fail_group.lean │ ├── fail_group.tptp │ ├── fail_lattice.lean │ ├── fail_lattice.tptp │ ├── fail_tests.lean │ ├── ffffa.lean │ ├── fluidSupTests.lean │ ├── group.lean │ ├── group2.lean │ ├── lastAsylum.lean │ ├── monobugs.lean │ ├── profiling.lean │ ├── temp.lean │ ├── temp2.lean │ ├── test_Order.lean │ ├── test_PUZ.lean │ ├── test_TPTP.lean │ ├── test_boolSimp.lean │ ├── test_continuity.lean │ ├── test_higher_order.lean │ ├── test_inhabitationReasoning.lean │ └── test_regression.lean ├── Unif.lean └── Util │ ├── AbstractMVars.lean │ ├── ClauseSubsumptionCheck.lean │ ├── Debugging.lean │ ├── DeeplyOccurringVars.lean │ ├── IdStrategyHeap.lean │ ├── InspectCoreM.lean │ ├── LazyList.lean │ ├── MessageData.lean │ ├── Misc.lean │ ├── OccursCheck.lean │ ├── ProofReconstruction.lean │ ├── Reduction.lean │ ├── RestoreCoreM.lean │ ├── RestoreCoreMCodeGen.lean │ ├── StrategyHeap.lean │ └── TypeInhabitationReasoning.lean ├── DuperOnMathlib ├── .gitignore ├── DuperOnMathlib.lean ├── DuperOnMathlib │ ├── CSB.lean │ ├── Cantor.lean │ ├── DuperInterface.lean │ ├── Function.lean │ ├── HigherOrder.lean │ ├── IVT.lean │ ├── ListExamples.lean │ ├── MIL │ │ ├── C02_Basics │ │ │ ├── S01_Calculating.lean │ │ │ ├── S02_Proving_Identities_in_Algebraic_Structures.lean │ │ │ ├── S03_Using_Theorems_and_Lemmas.lean │ │ │ ├── S04_More_on_Order_and_Divisibility.lean │ │ │ └── S05_Proving_Facts_about_Algebraic_Structures.lean │ │ └── C03_Logic │ │ │ ├── S01_Implication_and_the_Universal_Quantifier.lean │ │ │ ├── S02_The_Existential_Quantifier.lean │ │ │ ├── S03_Negation.lean │ │ │ ├── S04_Conjunction_and_Iff.lean │ │ │ ├── S05_Disjunction.lean │ │ │ └── S06_Sequences_and_Convergence.lean │ ├── Prime.lean │ ├── Prime2.lean │ ├── PrimesModFour.lean │ └── Set2.lean ├── lake-manifest.json ├── lakefile.lean ├── lakefile.olean └── lean-toolchain ├── LICENSE ├── Main.lean ├── README.md ├── TODO.md ├── lake-manifest.json ├── lakefile.lean ├── lean-toolchain └── starexec ├── .gitignore ├── bin └── starexec_run_default.sh └── pack_starexec.sh /.github/tff_tests.txt: -------------------------------------------------------------------------------- 1 | COM001_1 2 | COM002_1 3 | COM002_2 4 | COM003_1 5 | COM032_5 6 | COM033_5 7 | COM034_5 8 | COM035_5 9 | COM036_5 10 | COM037_5 11 | COM040_5 12 | COM041_5 13 | COM042_5 14 | COM043_5 15 | COM044_5 16 | COM045_5 17 | COM047_5 18 | COM069_5 19 | COM080_5 20 | COM088_5 21 | COM090_5 22 | COM092_5 23 | COM095_5 24 | COM111_5 25 | COM113_5 26 | COM122_5 27 | HWV039_3 28 | HWV040_3 29 | HWV041_3 30 | HWV042_1 31 | HWV042_3 32 | HWV043_3 33 | HWV044_3 34 | HWV045_3 35 | HWV046_3 36 | HWV047_3 37 | HWV048_1 38 | HWV048_3 39 | HWV049_1 40 | HWV050_3 41 | HWV051_3 42 | ITP001_1 43 | ITP001_2 44 | ITP001_3 45 | ITP001_5 46 | ITP001_7 47 | ITP004_3 48 | ITP004_7 49 | ITP006_3 50 | ITP010_1 51 | ITP010_2 52 | ITP010_3 53 | ITP010_7 54 | ITP011_3 55 | ITP012_3 56 | ITP012_7 57 | ITP013_3 58 | ITP015_3 59 | ITP015_7 60 | ITP018_1 61 | ITP018_3 62 | ITP018_7 63 | ITP019_1 64 | ITP019_2 65 | ITP019_3 66 | ITP019_7 67 | ITP020_3 68 | ITP021_3 69 | ITP021_7 70 | ITP023_1 71 | ITP023_3 72 | ITP209_1 73 | ITP209_2 74 | ITP222_2 75 | ITP227_2 76 | KRS003_1 77 | LCL748_5 78 | LCL755_5 79 | LCL756_5 80 | LCL759_5 81 | LCL760_5 82 | LCL762_5 83 | LCL763_5 84 | LCL764_5 85 | LCL769_5 86 | LCL770_5 87 | LCL771_5 88 | LCL773_5 89 | LCL774_5 90 | LCL775_5 91 | LCL776_5 92 | LCL777_5 93 | LCL778_5 94 | LCL779_5 95 | LCL784_5 96 | LCL785_5 97 | LCL786_5 98 | LCL787_5 99 | LCL788_5 100 | LCL789_5 101 | LCL792_5 102 | LCL793_5 103 | LCL795_5 104 | LCL796_5 105 | LCL798_5 106 | LCL799_5 107 | LCL800_5 108 | LCL802_5 109 | LCL803_5 110 | LCL805_5 111 | LCL806_5 112 | LCL807_5 113 | LCL808_5 114 | LCL809_5 115 | LCL811_5 116 | LCL812_5 117 | LCL813_5 118 | LCL814_5 119 | LCL815_5 120 | LCL816_5 121 | LCL817_5 122 | LCL818_5 123 | LCL820_5 124 | LCL821_5 125 | LCL822_5 126 | LCL823_5 127 | LCL825_5 128 | LCL826_5 129 | LCL827_5 130 | LCL828_5 131 | LCL829_5 132 | LCL830_5 133 | LCL834_5 134 | LCL835_5 135 | LCL836_5 136 | LCL838_5 137 | LCL839_5 138 | LCL840_5 139 | LCL843_5 140 | NUM924_2 141 | NUM924_3 142 | NUM925_1 143 | NUM925_2 144 | NUM925_3 145 | NUM925_5 146 | NUM931_5 147 | NUM940_5 148 | NUM942_5 149 | NUM944_5 150 | NUM945_5 151 | NUM948_5 152 | NUM952_5 153 | NUM954_5 154 | NUM957_5 155 | NUM958_5 156 | NUM960_5 157 | NUM962_5 158 | NUM967_5 159 | NUM970_5 160 | NUM971_5 161 | NUM972_5 162 | NUM973_5 163 | NUM977_5 164 | NUM978_5 165 | NUM980_5 166 | NUM983_5 167 | NUM991_5 168 | NUM994_5 169 | NUM996_5 170 | NUN004_5 171 | NUN008_5 172 | NUN015_5 173 | NUN016_5 174 | NUN017_5 175 | NUN018_5 176 | PHI044_1 177 | PUZ012_1 178 | PUZ018_1 179 | PUZ031_1 180 | PUZ130_1 181 | PUZ131_1 182 | PUZ134_2 183 | PUZ135_1 184 | PUZ135_2 185 | PUZ139_1 186 | SCT171_1 187 | SCT171_2 188 | SCT171_3 189 | SCT174_5 190 | SCT176_5 191 | SCT187_5 192 | SCT192_5 193 | SCT193_5 194 | SCT194_5 195 | SCT195_5 196 | SCT196_5 197 | SCT197_5 198 | SCT199_5 199 | SCT203_5 200 | SCT222_5 201 | SCT223_5 202 | SCT224_5 203 | SCT226_5 204 | SCT227_5 205 | SCT229_5 206 | SCT231_5 207 | SCT233_5 208 | SCT234_5 209 | SCT235_5 210 | SCT236_5 211 | SCT237_5 212 | SCT240_5 213 | SCT241_5 214 | SCT242_5 215 | SCT243_5 216 | SCT246_5 217 | SCT247_5 218 | SCT250_5 219 | SCT252_5 220 | SCT253_5 221 | SCT261_5 222 | SCT264_5 223 | SWV573_5 224 | SWV617_5 225 | SWV629_5 226 | SWV633_5 227 | SWV743_5 228 | SWV751_5 229 | SWV763_5 230 | SWV779_5 231 | SWV811_5 232 | SWW469_1 233 | SWW469_5 234 | SWW473_1 235 | SWW473_3 236 | SWW478_1 237 | SWW478_2 238 | SWW486_5 239 | SWW494_5 240 | SWW495_5 241 | SWW496_5 242 | SWW497_5 243 | SWW498_5 244 | SWW499_5 245 | SWW503_5 246 | SWW504_5 247 | SWW506_5 248 | SWW508_5 249 | SWW512_5 250 | SWW513_5 251 | SWW515_5 252 | SWW520_5 253 | SWW522_5 254 | SWW529_5 255 | SWW533_5 256 | SWW535_5 257 | SWW542_5 258 | SWW551_5 259 | SWW552_5 260 | SWW553_5 261 | SWW562_5 262 | SWW567_5 263 | SWW570_5 264 | SWW572_5 265 | -------------------------------------------------------------------------------- /.github/workflows/test.yml: -------------------------------------------------------------------------------- 1 | name: Test 2 | on: [push] 3 | jobs: 4 | Test: 5 | runs-on: ubuntu-latest 6 | steps: 7 | - name: Check out repository code 8 | uses: actions/checkout@v3 9 | 10 | - name: Install elan & Lean 11 | run: | 12 | set -o pipefail 13 | curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y 14 | ~/.elan/bin/lean --version 15 | echo "$HOME/.elan/bin" >> $GITHUB_PATH 16 | 17 | - name: Download TPTP 18 | run: | 19 | set -o pipefail 20 | wget https://www.tptp.org/TPTP/Archive/TPTP-v8.0.0.tgz 21 | tar -zxvf TPTP-v8.0.0.tgz -C .. 22 | 23 | - name: lake build 24 | run: | 25 | set -o pipefail 26 | lake build 27 | 28 | - name: Test THF 29 | run: | 30 | export LEAN_PATH=./build/lib 31 | while read p; do 32 | echo "THF Problem: $p" 33 | echo "import Duper.TPTP import Duper.Tactic set_option inhabitationReasoning false tptp ${p:0:6} \"../TPTP-v8.0.0/Problems/${p:0:3}/$p.p\" by duper_no_timing"\ 34 | | lake env ~/.elan/bin/lean -R ./. -o ./build/myfile.olean -i ./build/myfile.ilean --stdin --load-dynlib ./build/lib/libDuper-Tactic-1.so || true 35 | done <./.github/thf_tests.txt 36 | 37 | - name: Test TFF 38 | run: | 39 | export LEAN_PATH=./build/lib 40 | while read p; do 41 | echo "TFF Problem: $p" 42 | echo "import Duper.TPTP import Duper.Tactic set_option inhabitationReasoning false tptp ${p:0:6} \"../TPTP-v8.0.0/Problems/${p:0:3}/$p.p\" by duper_no_timing"\ 43 | | lake env ~/.elan/bin/lean -R ./. -o ./build/myfile.olean -i ./build/myfile.ilean --stdin --load-dynlib ./build/lib/libDuper-Tactic-1.so || true 44 | done <./.github/tff_tests.txt 45 | 46 | - name: Test FOF 47 | run: | 48 | export LEAN_PATH=./build/lib 49 | while read p; do 50 | echo "FOF Problem: $p" 51 | echo "import Duper.TPTP import Duper.Tactic set_option inhabitationReasoning false tptp ${p:0:6} \"../TPTP-v8.0.0/Problems/${p:0:3}/$p.p\" have inhabitedIota : Inhabited TPTP.iota := sorry; by duper_no_timing"\ 52 | | lake env ~/.elan/bin/lean -R ./. -o ./build/myfile.olean -i ./build/myfile.ilean --stdin --load-dynlib ./build/lib/libDuper-Tactic-1.so || true 53 | done <./.github/fof_tests.txt 54 | -------------------------------------------------------------------------------- /.github/workflows/test_exe.yml: -------------------------------------------------------------------------------- 1 | name: Test Executable 2 | on: [push] 3 | jobs: 4 | TestExe: 5 | runs-on: ubuntu-latest 6 | steps: 7 | - name: Check out repository code 8 | uses: actions/checkout@v3 9 | 10 | - name: Install elan & Lean 11 | run: | 12 | set -o pipefail 13 | curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y 14 | ~/.elan/bin/lean --version 15 | echo "$HOME/.elan/bin" >> $GITHUB_PATH 16 | 17 | - name: Download TPTP 18 | run: | 19 | set -o pipefail 20 | wget https://www.tptp.org/TPTP/Archive/TPTP-v8.0.0.tgz 21 | tar -zxvf TPTP-v8.0.0.tgz -C .. 22 | 23 | - name: lake build 24 | run: | 25 | set -o pipefail 26 | lake build 27 | 28 | - name: Test THF 29 | run: | 30 | while read p; do 31 | echo "THF Problem: $p" 32 | timeout 30 ./build/bin/duper "../TPTP-v8.0.0/Problems/${p:0:3}/$p.p" --github || true 33 | done <./.github/thf_tests.txt 34 | 35 | - name: Test TFF 36 | run: | 37 | while read p; do 38 | echo "TFF Problem: $p" 39 | timeout 30 ./build/bin/duper "../TPTP-v8.0.0/Problems/${p:0:3}/$p.p" --github || true 40 | done <./.github/tff_tests.txt 41 | 42 | - name: Test FOF 43 | run: | 44 | while read p; do 45 | echo "FOF Problem: $p" 46 | timeout 30 ./build/bin/duper "../TPTP-v8.0.0/Problems/${p:0:3}/$p.p" --github || true 47 | done <./.github/fof_tests.txt -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | .lake 2 | DuperOnMathlib/.lake 3 | /build 4 | /lake-packages 5 | /lean_packages 6 | /lakefile.olean 7 | .DS_Store 8 | -------------------------------------------------------------------------------- /Duper.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | -------------------------------------------------------------------------------- /Duper/BackwardSimplification.lean: -------------------------------------------------------------------------------- 1 | import Duper.ProverM 2 | import Duper.Simp 3 | import Duper.Rules.ClauseSubsumption 4 | import Duper.Rules.ContextualLiteralCutting 5 | import Duper.Rules.Demodulation 6 | import Duper.Rules.EqualitySubsumption 7 | import Duper.Rules.SimplifyReflect 8 | 9 | namespace Duper 10 | namespace ProverM 11 | 12 | def backwardSimpRules : ProverM (Array BackwardSimpRule) := do 13 | let subsumptionTrie ← getSubsumptionTrie 14 | return #[ 15 | (backwardDemodulation (← getDemodMainPremiseIdx)).toBackwardSimpRule, 16 | (backwardClauseSubsumption subsumptionTrie).toBackwardSimpRule, 17 | (backwardEqualitySubsumption subsumptionTrie).toBackwardSimpRule, 18 | (backwardContextualLiteralCutting subsumptionTrie).toBackwardSimpRule, 19 | (backwardPositiveSimplifyReflect subsumptionTrie).toBackwardSimpRule, 20 | (backwardNegativeSimplifyReflect subsumptionTrie).toBackwardSimpRule 21 | ] 22 | 23 | /-- Uses the givenClause to attempt to simplify other clauses in the active set. For each clause that backwardSimplify is 24 | able to produce a simplification for, backwardSimplify removes the clause adds any newly simplified clauses to the passive set. 25 | Additionally, for each clause removed from the active set in this way, all descendents of said clause should also be removed from 26 | the current state's allClauses and passiveSet -/ 27 | def backwardSimplify (givenClause : Clause) : ProverM Unit := do 28 | trace[duper.prover.saturate] "backward simplify with {givenClause}" 29 | let backwardSimpRules ← backwardSimpRules 30 | for i in [0 : backwardSimpRules.size] do 31 | let simpRule := backwardSimpRules[i]! 32 | simpRule givenClause 33 | -------------------------------------------------------------------------------- /Duper/DUnif/DApply.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import Duper.DUnif.UnifRules 3 | import Duper.Expr 4 | open Lean 5 | 6 | namespace DUnif 7 | 8 | -- Note: This is copied from standard library with some code 9 | -- removed to make it simpler and a few lines changed to 10 | -- allow for higher-order unification 11 | 12 | private def throwApplyError {α} (mvarId : MVarId) (eType : Expr) (targetType : Expr) : MetaM α := 13 | Meta.throwTacticEx `apply mvarId m!"failed to unify{indentExpr eType}\nwith{indentExpr targetType}" 14 | 15 | /-- 16 | Close the given goal using `apply e`. 17 | -/ 18 | def execdapply (mvarId : MVarId) (e : Expr) (nAttempt : Nat) (nUnif : Nat) (cont : Nat) (cfg : Meta.ApplyConfig := {}) : MetaM (List MVarId) := 19 | mvarId.withContext do 20 | mvarId.checkNotAssigned `apply 21 | let targetType ← mvarId.getType 22 | let targetMVars ← Meta.getMVarsNoDelayed targetType 23 | let eType ← Meta.inferType e 24 | let (numArgs, hasMVarHead) ← Meta.getExpectedNumArgsAux eType 25 | /- 26 | The `apply` tactic adds `_`s to `e`, and some of these `_`s become new goals. 27 | When `hasMVarHead` is `false` we try different numbers, until we find a type compatible with `targetType`. 28 | We used to try only `numArgs-targetTypeNumArgs` when `hasMVarHead = false`, but this is not always correct. 29 | For example, consider the following example 30 | ``` 31 | example {α β} [LE_trans β] (x y z : α → β) (h₀ : x ≤ y) (h₁ : y ≤ z) : x ≤ z := by 32 | apply le_trans 33 | assumption 34 | assumption 35 | ``` 36 | In this example, `targetTypeNumArgs = 1` because `LE` for functions is defined as 37 | ``` 38 | instance {α : Type u} {β : Type v} [LE β] : LE (α → β) where 39 | le f g := ∀ i, f i ≤ g i 40 | ``` 41 | -/ 42 | let rangeNumArgs ← if hasMVarHead then 43 | pure [numArgs : numArgs+1] 44 | else 45 | let targetTypeNumArgs ← Meta.getExpectedNumArgs targetType 46 | pure [numArgs - targetTypeNumArgs : numArgs+1] 47 | /- 48 | Auxiliary function for trying to add `n` underscores where `n ∈ [i: rangeNumArgs.stop)` 49 | See comment above 50 | -/ 51 | let rec go (i : Nat) : MetaM (Array Expr × Array BinderInfo) := do 52 | if i < rangeNumArgs.stop then 53 | let s ← saveState 54 | let (newMVars, binderInfos, eType) ← Meta.forallMetaTelescopeReducing eType i 55 | if (← hounif eType targetType nAttempt nUnif cont true) then 56 | return (newMVars, binderInfos) 57 | else 58 | s.restore 59 | go (i+1) 60 | else 61 | let (_, _, eType) ← Meta.forallMetaTelescopeReducing eType (some rangeNumArgs.start) 62 | throwApplyError mvarId eType targetType 63 | let (newMVars, binderInfos) ← go rangeNumArgs.start 64 | Meta.postprocessAppMVars `apply mvarId newMVars binderInfos cfg.synthAssignedInstances 65 | let e ← instantiateMVars e 66 | mvarId.assign (mkAppN e newMVars) 67 | let newMVars ← newMVars.filterM fun mvar => not <$> mvar.mvarId!.isAssigned 68 | -- Collect other mvars 69 | let mut otherMVarIds ← Meta.getMVarsNoDelayed e 70 | for m in targetMVars do 71 | for mvarId in (← Meta.getMVarsNoDelayed (.mvar m)) do 72 | if !otherMVarIds.contains mvarId then 73 | otherMVarIds := otherMVarIds.push mvarId 74 | let newMVarIds := (newMVars.map (·.mvarId!)).data 75 | otherMVarIds := otherMVarIds.filter fun mvarId => !newMVarIds.contains mvarId 76 | let result := newMVarIds ++ otherMVarIds.toList 77 | trace[Meta.Tactic] "{result}" 78 | result.forM (·.headBetaType) 79 | return result 80 | termination_by go i => rangeNumArgs.stop - i 81 | 82 | syntax (name := dapply) "dapply " term " attempt " num "unifier " num "contains" num : tactic 83 | 84 | @[tactic dapply] 85 | def evaldapply : Elab.Tactic.Tactic := fun stx => 86 | match stx with 87 | | `(tactic| dapply $e attempt $nAttempt unifier $nunif contains $cont) => Elab.Tactic.withMainContext do 88 | let mut val ← instantiateMVars (← Elab.Tactic.elabTermForApply e) 89 | if val.isMVar' then 90 | Elab.Term.synthesizeSyntheticMVarsNoPostponing 91 | val ← instantiateMVars val 92 | let mvarIds' ← execdapply (← Elab.Tactic.getMainGoal) val nAttempt.getNat nunif.getNat cont.getNat 93 | Elab.Term.synthesizeSyntheticMVarsNoPostponing 94 | Elab.Tactic.replaceMainGoal mvarIds' 95 | | _ => Elab.throwUnsupportedSyntax -------------------------------------------------------------------------------- /Duper/DUnif/DRefl.lean: -------------------------------------------------------------------------------- 1 | import Lean.Meta.Reduce 2 | import Lean.Meta.Tactic.Util 3 | import Lean.Meta.Tactic.Apply 4 | import Duper.DUnif.UnifRules 5 | 6 | open Lean Meta 7 | namespace DUnif 8 | 9 | register_option drefl.reduce : Bool := { 10 | defValue := False 11 | } 12 | 13 | initialize 14 | registerTraceClass `drefl.debug 15 | 16 | /-- 17 | Close given goal using `Eq.refl`. 18 | -/ 19 | def MVarId.refl (mvarId : MVarId) (nAttempt : Nat) (nUnif : Nat) (cont : Nat) (iterOn : Bool) : MetaM (Array MVarId) := do 20 | mvarId.withContext do 21 | mvarId.checkNotAssigned `refl 22 | let targetType ← mvarId.getType' 23 | unless targetType.isAppOfArity ``Eq 3 do 24 | throwTacticEx `rfl mvarId m!"equality expected{indentExpr targetType}" 25 | let mut lhs ← instantiateMVars targetType.appFn!.appArg! 26 | let mut rhs ← instantiateMVars targetType.appArg! 27 | if drefl.reduce.get (← getOptions) then 28 | let red (e : Expr) : MetaM TransformStep := do 29 | let e := e.consumeMData 30 | let e ← Meta.whnf e 31 | return .continue e 32 | lhs ← Meta.transform lhs (pre := red) 33 | rhs ← Meta.transform rhs (pre := red) 34 | trace[drefl.debug] "lhs: {lhs}, rhs: {rhs}" 35 | let success ← hounif lhs rhs nAttempt nUnif cont iterOn 36 | unless success do 37 | throwTacticEx `rfl mvarId m!"equality lhs{indentExpr lhs}\nis not definitionally equal to rhs{indentExpr rhs}" 38 | let us := targetType.getAppFn.constLevels! 39 | let α := targetType.appFn!.appFn!.appArg! 40 | mvarId.assign (mkApp2 (mkConst ``Eq.refl us) α lhs) 41 | Meta.getMVarsNoDelayed (.mvar mvarId) 42 | 43 | syntax (name := drefl) "drefl" " attempt " num "unifier " num "contains" num ("iteron")? : tactic 44 | 45 | @[tactic drefl] def evalRefl : Elab.Tactic.Tactic := fun stx => 46 | match stx with 47 | | `(tactic| drefl attempt $nAttempt unifier $nunif contains $cont iteron) => 48 | Elab.Tactic.liftMetaTactic fun mvarId => do 49 | let ids ← DUnif.MVarId.refl mvarId nAttempt.getNat nunif.getNat cont.getNat true 50 | return ids.data 51 | | `(tactic| drefl attempt $nAttempt unifier $nunif contains $cont) => 52 | Elab.Tactic.liftMetaTactic fun mvarId => do 53 | let ids ← DUnif.MVarId.refl mvarId nAttempt.getNat nunif.getNat cont.getNat false 54 | return ids.data 55 | | _ => Elab.throwUnsupportedSyntax 56 | 57 | end DUnif 58 | -------------------------------------------------------------------------------- /Duper/DUnif/Oracles.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import Duper.DUnif.Utils 3 | import Duper.DUnif.UnifProblem 4 | import Duper.Util.OccursCheck 5 | 6 | open Lean 7 | 8 | namespace DUnif 9 | 10 | initialize 11 | registerTraceClass `duper.dUnif.oracles 12 | 13 | register_option oracleInstOn : Bool := { 14 | defValue := true 15 | descr := "Whether to use instantiation oracle, which attempts to " ++ 16 | "instantiate a side of the equation if it's a metavariable" 17 | } 18 | 19 | def getOracleInstOn : CoreM Bool := do 20 | let opts ← getOptions 21 | return oracleInstOn.get opts 22 | 23 | def oracleInst (p : UnifProblem) (eq : UnifEq) : MetaM (Option UnifProblem) := do 24 | setMCtx p.mctx 25 | let mut mvarId : MVarId := Inhabited.default 26 | if ¬ (← getOracleInstOn) then 27 | return none 28 | let mut eq := eq 29 | if let .some id ← metaEta eq.rhs then 30 | eq := eq.swapSide 31 | mvarId := id 32 | else 33 | if let .some id ← metaEta eq.lhs.eta then 34 | mvarId := id 35 | else 36 | return none 37 | -- We do not use Lean's occurs check here, because the occurs check 38 | -- does not consider metavariables in the type of metavariables 39 | if (← mustNotOccursCheck mvarId eq.rhs) then 40 | mvarId.assign eq.rhs 41 | return some {(← p.pushParentRuleIfDbgOn (.OracleInst eq)) with checked := false, mctx := ← getMCtx} 42 | else 43 | return none 44 | 45 | register_option oracleOccursOn : Bool := { 46 | defValue := true 47 | descr := "Whether to use the occurs check oracle, which is the generalization" ++ 48 | " of first-order occurs check to dependent type theory. It only checks metavariables" ++ 49 | " which will always be in the term even after β-reduction and further metavariable instantiations." 50 | } 51 | 52 | def getOracleOccursOn : CoreM Bool := do 53 | let opts ← getOptions 54 | return oracleOccursOn.get opts 55 | 56 | -- true : fail 57 | -- false : not fail 58 | def oracleOccurs (p : UnifProblem) (eq : UnifEq) : MetaM Bool := do 59 | setMCtx p.mctx 60 | if ¬ (← getOracleOccursOn) then 61 | return false 62 | if let .mvar id := eq.rhs.eta then 63 | mustOccursCheck id eq.lhs 64 | else 65 | if let .mvar id := eq.lhs.eta then 66 | mustOccursCheck id eq.rhs 67 | else 68 | return false 69 | -------------------------------------------------------------------------------- /Duper/DUnif/Readme.md: -------------------------------------------------------------------------------- 1 | __Debugging__ 2 | * To debug unification procedure, use 3 | ``` 4 | set_option trace.duper.dUnif.debug true 5 | ``` 6 | * To enable advanced debugging functionality, use 7 | ``` 8 | set_option dUnifDbgOn true 9 | ``` 10 | * The ```contains ``` option of ```dapply``` and ```drefl``` is useless when ```dUnifDbgOn``` is ```false```, because we do not push parent clause and can't check ```contains``` when ```dUnifDbgOn``` is ```false```. -------------------------------------------------------------------------------- /Duper/DUnif/Utils.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | open Lean 3 | 4 | namespace DUnif 5 | 6 | def metaEta (e : Expr) : MetaM (Option MVarId) := do 7 | let mctx ← getMCtx 8 | let e ← instantiateMVars e 9 | let result ← metaEtaAux e 10 | match result with 11 | | .some (headId, F, hmap) => 12 | let headTy ← Meta.inferType (.mvar headId) 13 | let binding ← Meta.withDefault <| Meta.forallBoundedTelescope headTy hmap.size fun xs _ => do 14 | if hmap.size != xs.size then 15 | throwError "metaEta :: Unexpected error" 16 | let mut ret := F 17 | for argIdx in hmap do 18 | let .some x := xs[argIdx]? 19 | | throwError "metaEta :: Unexpected error" 20 | ret := Expr.app ret x 21 | Meta.mkLambdaFVars xs ret 22 | headId.assign binding 23 | let .mvar id := (← instantiateMVars F).eta 24 | | throwError "metaEta :: Unexpected error" 25 | return id 26 | | .none => setMCtx mctx; return none 27 | where metaEtaAux (e : Expr) : MetaM (Option (MVarId × Expr × Array Nat)) := 28 | Meta.lambdaTelescope e fun xs body => do 29 | let head := body.getAppFn 30 | let args := body.getAppArgs 31 | let .mvar headId := head 32 | | return none 33 | if args.size != xs.size then 34 | return none 35 | -- hmap : Map from position in `xs` to argIdx of `head` 36 | let xsMap := @Std.HashMap.ofList Expr Nat inferInstance inferInstance 37 | xs.zipIdx.toList 38 | let mut hmap : Array Nat := ⟨List.range xs.size⟩ 39 | let mut hset : Std.HashSet Nat := {} 40 | for (arg, argIdx) in args.zipIdx do 41 | let .some xsIdx := xsMap[arg]? 42 | | return .none 43 | if hset.contains xsIdx then return .none 44 | hset := hset.insert xsIdx 45 | hmap := hmap.setIfInBounds xsIdx argIdx 46 | let F ← Meta.mkFreshExprMVar (← Meta.inferType head) 47 | let F ← Meta.mkLambdaFVars xs F 48 | return .some (headId, F, hmap) 49 | 50 | /-- Make the most general function type with `n` explicit binders -/ 51 | def mkGeneralFnTy (n : Nat) (resTy : Option Expr := .none) : MetaM Expr := 52 | match n with 53 | | 0 => 54 | match resTy with 55 | | .some resTy => return resTy 56 | | .none => Meta.mkFreshTypeMVar 57 | | n + 1 => do 58 | let bty ← Meta.mkFreshTypeMVar 59 | let userName := Name.mkStr1 s!"gen{n}" 60 | Meta.withLocalDeclD userName bty fun x => do 61 | let res ← mkGeneralFnTy n resTy 62 | Meta.mkForallFVars #[x] <| res 63 | 64 | /-- Make an implication with `n` explicit binders -/ 65 | def mkImplication (n : Nat) (resTy : Option Expr := .none) : MetaM Expr := do 66 | let argTys_ty ← (Array.mk (List.range n)).mapM (fun _ => Meta.mkFreshLevelMVar) 67 | let argTys ← argTys_ty.mapM (fun e => Meta.mkFreshExprMVar (Expr.sort e)) 68 | let declInfos := argTys.zipIdx.map (fun (argTy, nat) => (Name.mkStr1 s!"imp{nat}", fun _ => pure argTy)) 69 | let resTy ← (do 70 | match resTy with 71 | | .some resTy => return resTy 72 | | .none => Meta.mkFreshTypeMVar) 73 | Meta.withLocalDeclsD declInfos fun xs => Meta.mkForallFVars xs resTy 74 | 75 | end DUnif 76 | -------------------------------------------------------------------------------- /Duper/Match.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import Duper.Util.Reduction 3 | 4 | set_option linter.unusedVariables false 5 | 6 | open Lean Lean.Meta Duper 7 | 8 | initialize Lean.registerTraceClass `duper.match.debug 9 | 10 | /-- Given an array of expression pairs (match_target, e), attempts to assign mvars in e to make e equal to match_target (without 11 | making any assignments to mvars that appear in protected_mvars). 12 | Returns true and performs mvar assignments if successful, returns false and does not perform any mvar assignments otherwise -/ 13 | partial def Lean.Meta.performMatch (l : Array (Expr × Expr)) (protected_mvars : Array MVarId) : MetaM Bool := do 14 | Core.checkMaxHeartbeats "match" 15 | let state ← saveState 16 | try 17 | trace[duper.match.debug] "About to attempt to match {l}" 18 | 19 | for (match_target, e) in l do 20 | let match_target_type ← betaEtaReduceInstMVars $ ← inferType match_target 21 | let e_type ← betaEtaReduceInstMVars $ ← inferType e 22 | if not (← match1 match_target_type e_type protected_mvars) then 23 | state.restore 24 | return false 25 | else if ← match1 match_target e protected_mvars then 26 | continue 27 | else 28 | state.restore 29 | return false 30 | 31 | trace[duper.match.debug] "Successfully matched {l}" 32 | return true 33 | catch ex => 34 | trace[duper.match.debug] "Encountered error {ex.toMessageData} in match procedure" 35 | state.restore 36 | throw ex 37 | where 38 | match1 (match_target e : Expr) (protected_mvars : Array MVarId) : MetaM Bool := do 39 | let match_target ← instantiateMVars match_target 40 | let e ← instantiateMVars e 41 | if match_target == e then return true else 42 | match_target.withApp fun match_target_hd match_target_tl => e.withApp fun e_hd e_tl => do 43 | match match_target_hd, e_hd with 44 | | Expr.fvar fid, Expr.fvar gid => do 45 | if fid == gid then 46 | matchArgs match_target_tl e_tl 47 | else 48 | return false 49 | | Expr.const fname flvls, Expr.const gname glvls => do 50 | if fname == gname && 51 | (← (flvls.zip glvls).allM (fun (flvl, glvl) => isLevelDefEq flvl glvl)) then 52 | matchArgs match_target_tl e_tl 53 | else 54 | return false 55 | | Expr.forallE .., Expr.forallE .. => 56 | return match_target_hd == e_hd && (← matchArgs match_target_tl e_tl) 57 | | Expr.lam .., Expr.lam .. => 58 | return match_target_hd == e_hd && (← matchArgs match_target_tl e_tl) 59 | | Expr.fvar .., Expr.mvar .. => 60 | matchRigidFlex match_target e protected_mvars 61 | | Expr.const .., Expr.mvar .. => 62 | matchRigidFlex match_target e protected_mvars 63 | | Expr.forallE .., Expr.mvar .. => 64 | matchRigidFlex match_target e protected_mvars 65 | | Expr.mvar .., Expr.mvar .. => 66 | matchFlexFlex match_target e protected_mvars 67 | | _, _ => return false 68 | matchArgs (match_target_tl e_tl : Array Expr) : MetaM Bool := do 69 | if match_target_tl.size == e_tl.size then 70 | performMatch (match_target_tl.zip e_tl) protected_mvars 71 | else 72 | return false 73 | matchRigidFlex (match_target e : Expr) (protected_mvars : Array MVarId) : MetaM Bool := do 74 | match_target.withApp fun match_target_hd match_target_tl => e.withApp fun e_hd e_tl => do 75 | match ← getExprMVarAssignment? e_hd.mvarId! with 76 | | some exp => match1 match_target (mkAppN exp e_tl) protected_mvars 77 | | none => do 78 | if e_tl.isEmpty then 79 | if (← occursCheck e_hd.mvarId! match_target) && not (protected_mvars.contains (e_hd.mvarId!)) then 80 | Lean.MVarId.assign e_hd.mvarId! match_target 81 | return true 82 | else return false 83 | else return false 84 | matchFlexFlex (match_target e : Expr) (protected_mvars : Array MVarId) : MetaM Bool := do 85 | match_target.withApp fun match_target_hd match_target_tl => e.withApp fun e_hd e_tl => do 86 | if match_target_tl.isEmpty && e_tl.isEmpty then 87 | if match_target_hd == e_hd then 88 | return true 89 | else 90 | if not (protected_mvars.contains (e_hd.mvarId!)) then 91 | Lean.MVarId.assign e_hd.mvarId! match_target 92 | return true 93 | else return false 94 | else return false 95 | -------------------------------------------------------------------------------- /Duper/Rules/ArgumentCongruence.lean: -------------------------------------------------------------------------------- 1 | import Duper.RuleM 2 | import Duper.MClause 3 | import Duper.Util.ProofReconstruction 4 | 5 | set_option linter.unusedVariables false 6 | 7 | namespace Duper 8 | open RuleM 9 | open Lean 10 | open Meta 11 | 12 | initialize Lean.registerTraceClass `duper.rule.argCong 13 | 14 | def mkArgumentCongruenceProof (i : Nat) (premises : List Expr) (parents : List ProofParent) 15 | (transferExprs : Array Expr) (c : Clause) : MetaM Expr := 16 | Meta.forallTelescope c.toForallExpr fun xs body => do 17 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 18 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 19 | let parentLits := parentsLits[0]! 20 | let appliedPremise := appliedPremises[0]! 21 | let mVarTys := transferExprs 22 | 23 | let mut caseProofs := Array.mkEmpty parentLits.size 24 | for j in [:parentLits.size] do 25 | let lit := parentLits[j]! 26 | if j == i then 27 | let newMVars ← (mVarTys.map some).mapM Meta.mkFreshExprMVar 28 | let newFVars ← newMVars.mapM Lean.instantiateMVars 29 | let cj := cLits[j]! 30 | let pj := lit 31 | let cjl := cj.lhs 32 | let pjl ← Meta.mkAppM' pj.lhs newMVars 33 | let able_to_match ← Meta.performMatch #[(cjl, pjl)] #[] 34 | if able_to_match then 35 | trace[duper.rule.argCong] m!"lhs of conclusion: {cjl}, lhs of parent: {pjl}" 36 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 37 | let mut pr := h 38 | for x in newFVars do 39 | pr ← Meta.mkAppM ``congrFun #[pr, x] 40 | trace[duper.rule.argCong] m!"pr: {pr}, type: {← Meta.inferType pr}" 41 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) j pr 42 | caseProofs := caseProofs.push pr 43 | else 44 | throwError s!"mkArgumentCongruenceProof:: " ++ 45 | s!"Unable to unify lhs of conclusion: {cjl} with " ++ 46 | s!"lhs of parent: {pjl}" 47 | else 48 | -- need proof of `L_j → L_1 ∨ ... ∨ L_n` 49 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 50 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) j h 51 | caseProofs := caseProofs.push pr 52 | 53 | let r ← orCases (parentLits.map Lit.toExpr) caseProofs 54 | Meta.mkLambdaFVars xs $ mkApp r appliedPremise 55 | 56 | def argCongAtLit (given : Clause) (c : MClause) (i : Nat) : RuleM (Array ClauseStream) := 57 | withoutModifyingMCtx $ do 58 | let lit := c.lits[i]! 59 | let mut streams := #[] 60 | if ← eligibilityNoUnificationCheck c (alreadyReduced := true) i (strict := true) then 61 | let ty ← inferType lit.lhs 62 | let (mVars, _, _) ← forallMetaTelescope ty 63 | trace[duper.rule.argCong] s!"Lhs: {lit.lhs}, Level: {lit.lvl}, Type of lhs: {ty}, Telescope: {mVars}" 64 | let lhs := lit.lhs; let rhs := lit.rhs; 65 | let mut newMVars := #[] 66 | let mut mVarTys := #[] 67 | let loaded ← getLoadedClauses 68 | for m in mVars do 69 | newMVars := newMVars.push m 70 | mVarTys := mVarTys.push (← instantiateMVars (← inferType m)) 71 | let newlhs := mkAppN lhs newMVars 72 | let newrhs := mkAppN rhs newMVars 73 | let newty ← inferType newlhs 74 | let newsort ← inferType newty 75 | let newlit := { lit with lhs := newlhs, 76 | rhs := newrhs, 77 | ty := ← inferType newlhs 78 | lvl := Expr.sortLevel! newsort} 79 | let c' := c.replaceLit! i newlit 80 | let ug ← unifierGenerator #[] 81 | let yC := do 82 | setLoadedClauses loaded 83 | yieldClause c' "argument congruence" (mkProof := mkArgumentCongruenceProof i) (transferExprs := mVarTys) 84 | streams := streams.push (ClauseStream.mk ug given yC "argument congruence") 85 | return streams 86 | 87 | def argCong (given : Clause) (c : MClause) (cNum : Nat) : RuleM (Array ClauseStream) := do 88 | trace[duper.rule.argCong] "ArgCong inferences with {c.lits}" 89 | let mut streams := #[] 90 | for i in [:c.lits.size] do 91 | if c.lits[i]!.sign = true then 92 | let str ← argCongAtLit given c i 93 | streams := streams.append str 94 | return streams 95 | 96 | end Duper 97 | -------------------------------------------------------------------------------- /Duper/Rules/BetaEtaReduction.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | import Duper.Util.ProofReconstruction 3 | import Duper.Util.Reduction 4 | 5 | set_option linter.unusedVariables false 6 | 7 | namespace Duper 8 | open Std 9 | open RuleM 10 | open SimpResult 11 | open Lean 12 | open Meta 13 | open Core 14 | 15 | initialize registerTraceClass `duper.rule.betaEtaReduce 16 | 17 | def mkBetaEtaReductionProof (premises : List Expr) (parents: List ProofParent) (transferExprs : Array Expr) 18 | (c : Clause) : MetaM Expr := do 19 | Meta.forallTelescope c.toForallExpr fun xs body => do 20 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 21 | let appliedPremise := appliedPremises[0]! 22 | Meta.mkLambdaFVars xs appliedPremise 23 | 24 | /-- Exhaustively apply beta, eta, and zeta reduction to every subterm in the clause -/ 25 | def betaEtaReduction : MSimpRule := fun c => do 26 | let c ← loadClause c 27 | let reducedC ← c.mapM (fun e => betaEtaReduce e) 28 | if reducedC == c then 29 | return none 30 | else 31 | let yc ← yieldClause reducedC "betaEtaReduce" $ some mkBetaEtaReductionProof 32 | return some #[yc] 33 | -------------------------------------------------------------------------------- /Duper/Rules/BoolHoist.lean: -------------------------------------------------------------------------------- 1 | import Duper.MClause 2 | import Duper.RuleM 3 | import Duper.Simp 4 | import Duper.Util.ProofReconstruction 5 | import Duper.Rules.IdentBoolHoist 6 | 7 | set_option linter.unusedVariables false 8 | 9 | namespace Duper 10 | open Lean 11 | open RuleM 12 | open Meta 13 | open SimpResult 14 | 15 | initialize Lean.registerTraceClass `duper.rule.boolHoist 16 | 17 | def boolHoistAtExpr (e : Expr) (pos : ClausePos) (given : Clause) (c : MClause) : RuleM (Array ClauseStream) := 18 | withoutModifyingMCtx do 19 | if c.lits[pos.lit]!.sign && pos.pos == #[] then -- e cannot be at the top level of a positive literal 20 | return #[] 21 | if (e.getTopSymbol).isMVar' then -- e cannot be variable headed 22 | return #[] 23 | if (e.isFullyAppliedLogicalSymbol) then -- e cannot be a fully applied logical symbol 24 | return #[] 25 | if not (← eligibilityNoUnificationCheck c (alreadyReduced := true) pos.lit) then 26 | -- No unificaiton check rather than PreUnification check because condition 3 talks about the position being eligible in 27 | -- the original clause (as opposed to being eligible in the clause with respect to the substitution σ) 28 | return #[] 29 | let eType ← inferType e 30 | let loaded ← getLoadedClauses 31 | let ug ← unifierGenerator #[(eType, .sort levelZero)] 32 | let yC := do 33 | setLoadedClauses loaded 34 | let lit := c.lits[pos.lit]! 35 | let eSide := lit.getSide pos.side 36 | let otherSide := lit.getOtherSide pos.side 37 | let cmp ← compare eSide otherSide false 38 | if cmp == Comparison.LessThan || cmp == Comparison.Equal then -- If eSide ≤ otherSide then e is not in an eligible position 39 | return none 40 | -- All side conditions have been met. Yield the appropriate clause 41 | let cErased := c.eraseLit pos.lit 42 | let some newLit ← lit.replaceAtPosUpdateType? ⟨pos.side, pos.pos⟩ (mkConst ``False) 43 | | return none -- If `lit` can't be safely changed at position `p`, then don't apply `boolHoist` at `p` 44 | let newClause := cErased.appendLits #[newLit, Lit.fromSingleExpr e true] 45 | trace[duper.rule.boolHoist] "Created {newClause.lits} from {c.lits}" 46 | yieldClause newClause "boolHoist" $ some (mkBoolHoistProof pos false) 47 | return #[ClauseStream.mk ug given yC "boolHoist"] 48 | 49 | def boolHoist (given : Clause) (c : MClause) (cNum : Nat) : RuleM (Array ClauseStream) := do 50 | trace[duper.rule.boolHoist] "Running BoolHoist on {c.lits}" 51 | let fold_fn := fun streams e pos => do 52 | let str ← boolHoistAtExpr e pos given c 53 | return streams.append str 54 | c.foldGreenM fold_fn #[] 55 | -------------------------------------------------------------------------------- /Duper/Rules/ClauseSubsumption.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | import Duper.Util.ClauseSubsumptionCheck 3 | 4 | namespace Duper 5 | 6 | open Lean Meta RuleM SimpResult 7 | 8 | initialize Lean.registerTraceClass `duper.rule.clauseSubsumption 9 | 10 | /-- Returns true if there exists a clause that subsumes c, and returns false otherwise -/ 11 | def forwardClauseSubsumption (subsumptionTrie : SubsumptionTrie) : MSimpRule := fun c => do 12 | let potentialSubsumingClauses ← subsumptionTrie.getPotentialSubsumingClauses c 13 | trace[duper.rule.clauseSubsumption] "number of potentialSubsumingClauses for {c}: {potentialSubsumingClauses.size}" 14 | let (cMVars, c) ← loadClauseCore c 15 | let cMVarIds := cMVars.map Expr.mvarId! 16 | let fold_fn := fun acc nextClause => do 17 | match acc with 18 | | false => 19 | conditionallyModifyingLoadedClauses do 20 | let nextClause ← loadClause nextClause 21 | if ← subsumptionCheck nextClause c cMVarIds then 22 | trace[duper.rule.clauseSubsumption] "Forward subsumption: removed {c.lits} because it was subsumed by {nextClause.lits}" 23 | return (true, true) 24 | else return (false, false) 25 | | true => return true 26 | let isSubsumed ← potentialSubsumingClauses.foldlM fold_fn false 27 | if isSubsumed then 28 | return some #[] 29 | else 30 | return none 31 | 32 | /-- Returns the list of clauses that givenSubsumingClause subsumes -/ 33 | def backwardClauseSubsumption (subsumptionTrie : SubsumptionTrie) : BackwardMSimpRule := fun givenSubsumingClause => do 34 | let potentialSubsumedClauses ← subsumptionTrie.getPotentialSubsumedClauses givenSubsumingClause 35 | trace[duper.rule.clauseSubsumption] "number potentialSubsumedClauses for {givenSubsumingClause}: {potentialSubsumedClauses.size}" 36 | let givenSubsumingClause ← loadClause givenSubsumingClause 37 | let fold_fn := fun acc nextClause => 38 | withoutModifyingMCtx do 39 | conditionallyModifyingLoadedClauses do 40 | let (nextClauseMVars, nextClauseM) ← loadClauseCore nextClause 41 | let nextClauseMVarIds := nextClauseMVars.map Expr.mvarId! 42 | if ← subsumptionCheck givenSubsumingClause nextClauseM nextClauseMVarIds then 43 | trace[duper.rule.clauseSubsumption] "Backward subsumption: removed {nextClause.lits} because it was subsumed by {givenSubsumingClause.lits}" 44 | return (true, acc.push nextClause) 45 | else return (false, acc) 46 | let clausesToRemove ← potentialSubsumedClauses.foldlM fold_fn #[] 47 | return clausesToRemove.map (fun x => (x, none)) 48 | -------------------------------------------------------------------------------- /Duper/Rules/DatatypeAcyclicity.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | import Duper.Util.ProofReconstruction 3 | 4 | set_option linter.unusedVariables false 5 | 6 | namespace Duper 7 | open Std 8 | open RuleM 9 | open SimpResult 10 | open Lean 11 | open Meta 12 | open LitSide 13 | 14 | initialize Lean.registerTraceClass `duper.rule.datatypeAcyclicity 15 | 16 | /-- Produces a list of (possibly duplicate) constructor subterms for `e` -/ 17 | partial def collectConstructorSubterms (e : Expr) : MetaM (Array Expr) := do 18 | let isConstructor ← matchConstCtor e.getAppFn' (fun _ => pure false) (fun _ _ => pure true) 19 | if isConstructor then 20 | let constructorSubterms ← e.getAppArgs.mapM (fun arg => collectConstructorSubterms arg) 21 | return constructorSubterms.flatten.push e 22 | else 23 | return #[e] 24 | 25 | /-- Returns `none` if `lit` does not compare constructor subterms, and returns `some litside` if `lit.litside` 26 | is a subterm of the constructor it is being compared to. Note that `lit.litside` may not itself be a constructor 27 | (e.g. `xs` is a constructor subterm of `x :: xs`) -/ 28 | def litComparesConstructorSubterms (lit : Lit) : MetaM (Option LitSide) := do 29 | let litTyIsInductive ← matchConstInduct lit.ty.getAppFn' (fun _ => pure false) (fun _ _ => pure true) 30 | if litTyIsInductive then 31 | trace[duper.rule.datatypeAcyclicity] "lit.ty {lit.ty} is an inductive datatype" 32 | -- If `e1` is a constructor subterm of `e2`, then `e1.weight ≤ e2.weight` 33 | if lit.lhs.weight < lit.rhs.weight then 34 | let rhsConstructorSubterms ← collectConstructorSubterms lit.rhs 35 | if rhsConstructorSubterms.contains lit.lhs then return some lhs 36 | else return none 37 | else if lit.rhs.weight < lit.lhs.weight then 38 | let lhsConstructorSubterms ← collectConstructorSubterms lit.lhs 39 | if lhsConstructorSubterms.contains lit.rhs then return some rhs 40 | else return none 41 | else 42 | if lit.lhs == lit.rhs then return some lhs 43 | else return none 44 | else -- `lit.ty` is not an inductive datatype so `lit` cannot be comparing constructor subterms 45 | trace[duper.rule.datatypeAcyclicity] "lit.ty {lit.ty} is not an inductive datatype" 46 | return none 47 | 48 | def mkDatatypeAcyclicityProof (removedLitNum : Nat) (litSide : LitSide) (premises : List Expr) 49 | (parents : List ProofParent) (transferExprs : Array Expr) (c : Clause) : MetaM Expr := do 50 | Meta.forallTelescope c.toForallExpr fun xs body => do 51 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 52 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 53 | let parentLits := parentsLits[0]! 54 | let appliedPremise := appliedPremises[0]! 55 | let mut proofCases : Array Expr := Array.mkEmpty parentLits.size 56 | for i in [:parentLits.size] do 57 | let lit := parentLits[i]! 58 | if i == removedLitNum then -- `lit` is the equality asserting an acyclic constructor 59 | let proofCase ← Meta.withLocalDeclD `h lit.toExpr fun h => do 60 | let sizeOfInst ← mkAppOptM ``inferInstanceAs #[← mkAppOptM ``SizeOf #[lit.ty], none] 61 | let litTyMVar ← mkFreshExprMVar lit.ty 62 | let abstrLam ← mkLambdaFVars #[litTyMVar] $ ← mkAppOptM ``sizeOf #[some lit.ty, some sizeOfInst, some litTyMVar] 63 | let sizeOfEq ← mkAppM ``congrArg #[abstrLam, h] -- Has the type `sizeOf lit.lhs = sizeOf lit.rhs` 64 | -- Need to generate a term of type `¬(sizeOf lit.lhs = sizeOf lit.rhs)` 65 | let sizeOfEqFalseMVar ← mkFreshExprMVar $ ← mkAppM ``Not #[← inferType sizeOfEq] -- Has the type `¬(sizeOf lit.lhs = sizeOf lit.rhs)` 66 | let sizeOfEqFalseMVarId := sizeOfEqFalseMVar.mvarId! 67 | -- **TODO**: Figure out how to assign `sizeOfEqFalseMVar` an actual term 68 | let proofCase := mkApp2 (mkConst ``False.elim [levelZero]) body $ mkApp sizeOfEqFalseMVar sizeOfEq -- Has the type `body` 69 | trace[duper.rule.datatypeAcyclicity] "lit: {lit}, lit.ty: {lit.ty}, sizeOfInst: {sizeOfInst}, abstrLam: {abstrLam}, sizeOfEq: {sizeOfEq}" 70 | trace[duper.rule.datatypeAcyclicity] "sizeOfEqFalseMVar: {sizeOfEqFalseMVar}, proofCase: {proofCase}" 71 | Meta.mkLambdaFVars #[h] proofCase 72 | proofCases := proofCases.push proofCase 73 | else -- `lit` is not the equality to be removed 74 | let proofCase ← Meta.withLocalDeclD `h lit.toExpr fun h => do 75 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) i h 76 | proofCases := proofCases.push proofCase 77 | let proof ← orCases (parentLits.map Lit.toExpr) proofCases 78 | Meta.mkLambdaFVars xs $ mkApp proof appliedPremise 79 | 80 | /-- Implements the acyclicity rules described in section 6.4 of https://arxiv.org/pdf/1611.02908 -/ 81 | def datatypeAcyclicity : MSimpRule := fun c => do 82 | let c ← loadClause c 83 | for i in [:c.lits.size] do 84 | let lit := c.lits[i]! 85 | match ← litComparesConstructorSubterms lit with 86 | | some side => 87 | if lit.sign then -- `lit` is never true so `lit` can be removed from `c` 88 | let res := c.eraseIdx i 89 | let yC ← yieldClause res "datatypeAcyclicity" none -- $ mkDatatypeAcyclicityProof i side 90 | trace[duper.rule.datatypeAcyclicity] "datatypeAcyclicity applied to {c.lits} to yield {yC.1}" 91 | return some #[yC] 92 | else -- `lit` is a tautology so the clause `c` can simply be removed 93 | trace[duper.rule.datatypeAcyclicity] "datatypeAcyclicity applied to remove {c.lits}" 94 | return some #[] 95 | | none => continue 96 | return none 97 | -------------------------------------------------------------------------------- /Duper/Rules/DecElim.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | import Duper.Util.ProofReconstruction 3 | import Lean.Meta.Basic 4 | 5 | set_option linter.unusedVariables false 6 | 7 | namespace Duper 8 | open RuleM 9 | open SimpResult 10 | open Lean 11 | open Meta 12 | 13 | initialize Lean.registerTraceClass `duper.rule.decElim 14 | 15 | /-- Checks whether a literal is decidably false. Returns true if the literal is decidably false and false otherwise. -/ 16 | def isDecidablyFalse (lit : Lit) : MetaM Bool := do 17 | try 18 | let d ← mkDecide lit.toExpr 19 | let d ← instantiateMVars d 20 | let r ← withDefault <| whnf d 21 | return r.isConstOf ``false 22 | catch _ => 23 | return false 24 | 25 | def mkDecElimProof (refs : List (Option Nat)) (premises : List Expr) (parents : List ProofParent) 26 | (transferExprs : Array Expr) (c : Clause) : MetaM Expr := 27 | Meta.forallTelescope c.toForallExpr fun xs body => do 28 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 29 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 30 | let parentLits := parentsLits[0]! 31 | let appliedPremise := appliedPremises[0]! 32 | 33 | if refs.length != parentLits.size then throwError "mkDecElimProof error: {refs} and {parentLits} have different sizes" 34 | 35 | let mut caseProofs := Array.mkEmpty parentLits.size 36 | for i in [:parentLits.size] do 37 | let lit := parentLits[i]! 38 | match refs[i]! with 39 | | none => 40 | -- This is adapted from the internals of `decide` 41 | let expectedType := lit.toExpr 42 | trace[duper.rule.decElim] "Trying to decide {expectedType}" 43 | let d ← mkDecide expectedType 44 | let d ← instantiateMVars d 45 | let r ← withDefault <| whnf d 46 | unless r.isConstOf ``false do 47 | throwError "mkDecElimProof: Failed to reduce to 'false'{indentExpr r}" 48 | trace[duper.rule.decElim] "{d} is false" 49 | let s := d.appArg! -- get instance from `d` 50 | let rflPrf ← mkEqRefl (toExpr false) 51 | let proofCase := mkApp3 (Lean.mkConst ``of_decide_eq_false) expectedType s rflPrf 52 | trace[duper.rule.decElim] "Built {proofCase} proving {d} is false" 53 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 54 | let proofCase := mkApp proofCase h 55 | let proofCase := mkApp2 (mkConst ``False.elim [levelZero]) body proofCase 56 | Meta.mkLambdaFVars #[h] proofCase 57 | caseProofs := caseProofs.push pr 58 | | some j => 59 | -- need proof of `L_i → L_1 ∨ ... ∨ L_n` 60 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 61 | let idx := j 62 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) idx h 63 | caseProofs := caseProofs.push pr 64 | let r ← orCases (parentLits.map Lit.toExpr) caseProofs 65 | Meta.mkLambdaFVars xs $ mkApp r appliedPremise 66 | 67 | /-- If there are any literals in `c` that are decidably true, then `decElim` removes `c`. Otherwise, if there are 68 | any literals in `c` that are decidably false, then `decElim` yields the clause obtained by removing all such literals. -/ 69 | def decElim : MSimpRule := fun c => do 70 | let c ← loadClause c 71 | /- 72 | Spec for newLits and refs: 73 | If c.lits[i] is decidably false, then refs[i] = none 74 | If c.lits[i] isn't decidably false, then refs[i] = some j where newLits[j] = c.lits[i] 75 | -/ 76 | let mut newLits : List Lit := [] 77 | let mut refs : List (Option Nat) := [] 78 | for lit in c.lits do 79 | if ← isDecidablyFalse lit then -- Remove the decidably false literal from the clause 80 | refs := none :: refs 81 | else 82 | refs := (some newLits.length) :: refs 83 | newLits := lit :: newLits 84 | -- To achieve the desired spec for newLits and refs, I must reverse them 85 | newLits := newLits.reverse 86 | refs := refs.reverse 87 | trace[duper.rule.decElim] "newLits: {newLits}, refs: {refs}" 88 | if (newLits.length = c.lits.size) then 89 | return none 90 | else 91 | let resultClause ← yieldClause (MClause.mk newLits.toArray) "decidable false elimination" 92 | (some (mkDecElimProof refs)) 93 | return some #[resultClause] 94 | 95 | end Duper 96 | -------------------------------------------------------------------------------- /Duper/Rules/DestructiveEqualityResolution.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | import Duper.Util.ProofReconstruction 3 | 4 | set_option linter.unusedVariables false 5 | 6 | namespace Duper 7 | open RuleM 8 | open SimpResult 9 | open Lean 10 | open Meta 11 | 12 | def is_var (e : Expr) : Bool := 13 | match Expr.consumeMData e with 14 | | Expr.mvar _ => true 15 | | Expr.fvar _ => true 16 | | Expr.bvar _ => true 17 | | _ => false 18 | 19 | def mkDestructiveEqualtiyResolutionProof (i : Nat) (premises : List Expr) (parents: List ProofParent) (transferExprs : Array Expr) 20 | (c : Clause) : MetaM Expr := do 21 | Meta.forallTelescope c.toForallExpr fun xs body => do 22 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 23 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 24 | let parentLits := parentsLits[0]! 25 | let appliedPremise := appliedPremises[0]! 26 | 27 | let mut caseProofs := Array.mkEmpty parentLits.size 28 | for j in [:parentLits.size] do 29 | let lit := parentLits[j]! 30 | if j == i then 31 | -- lit has the form t ≠ t 32 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 33 | let pr := mkApp2 (mkConst ``rfl [lit.lvl]) lit.ty lit.lhs 34 | let pr := mkApp h pr 35 | let pr := mkApp2 (mkConst ``False.elim [levelZero]) body pr 36 | Meta.mkLambdaFVars #[h] pr 37 | caseProofs := caseProofs.push pr 38 | else 39 | -- need proof of `L_j → L_1 ∨ ... ∨ L_n` 40 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 41 | let idx := if j ≥ i then j - 1 else j 42 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) idx h 43 | caseProofs := caseProofs.push $ pr 44 | 45 | let r ← orCases (parentLits.map Lit.toExpr) caseProofs 46 | Meta.mkLambdaFVars xs $ mkApp r appliedPremise 47 | 48 | def destructiveEqualityResolutionAtLit (c : MClause) (i : Nat) : RuleM (Option (Clause × Proof)) := 49 | withoutModifyingMCtx do 50 | let lit := c.lits[i]! 51 | if ← fastUnify #[(lit.lhs, lit.rhs)] then 52 | some <$> yieldClause (c.eraseLit i) "destructive equality resolution" 53 | (some (mkDestructiveEqualtiyResolutionProof i)) 54 | else 55 | return none -- Cannot apply destructive equality resolution to this literal, 56 | -- but it may still be possible to apply it to a different literal in the clause 57 | 58 | def destructiveEqualityResolution : MSimpRule := fun c => do 59 | let c ← loadClause c 60 | for i in [:c.lits.size] do 61 | if c.lits[i]!.sign = false ∧ (is_var c.lits[i]!.lhs ∨ is_var c.lits[i]!.rhs) then 62 | match ← destructiveEqualityResolutionAtLit c i with 63 | | some cp => return some #[cp] 64 | | none => continue 65 | return none 66 | -------------------------------------------------------------------------------- /Duper/Rules/ElimDupLit.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | import Duper.Util.ProofReconstruction 3 | 4 | set_option linter.unusedVariables false 5 | 6 | namespace Duper 7 | open RuleM 8 | open SimpResult 9 | open Lean 10 | open Meta 11 | 12 | initialize Lean.registerTraceClass `duper.rule.elimDupLit 13 | 14 | def mkElimDupLitProof (refs : Array (Nat × Bool)) (premises : List Expr) (parents: List ProofParent) (transferExprs : Array Expr) 15 | (c : Clause) : MetaM Expr := do 16 | Meta.forallTelescope c.toForallExpr fun xs body => do 17 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 18 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 19 | let parentLits := parentsLits[0]! 20 | let appliedPremise := appliedPremises[0]! 21 | 22 | let mut proofCases : Array Expr := Array.mkEmpty parentLits.size 23 | for i in [:parentLits.size] do 24 | let proofCase ← Meta.withLocalDeclD `h parentLits[i]!.toExpr fun h => do 25 | let (refsI, isFlipped) := refs[i]! 26 | if isFlipped then 27 | let appropriateSymmRule := 28 | if (parentLits[i]!.sign) then ``Eq.symm 29 | else ``Ne.symm 30 | let hSymm ← Meta.mkAppM appropriateSymmRule #[h] 31 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) refsI hSymm 32 | else 33 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) refsI h 34 | proofCases := proofCases.push proofCase 35 | let proof ← orCases (parentLits.map Lit.toExpr) proofCases 36 | let proof ← Meta.mkLambdaFVars xs $ mkApp proof appliedPremise 37 | return proof 38 | 39 | /-- Remove duplicate literals -/ 40 | def elimDupLit : MSimpRule := fun c => do 41 | let c ← loadClause c 42 | let mut newLits : Array Lit := Array.mkEmpty c.lits.size 43 | let mut refs : Array (Nat × Bool) := Array.mkEmpty c.lits.size 44 | for i in [:c.lits.size] do 45 | match newLits.finIdxOf? c.lits[i]! with 46 | | some j => do 47 | trace[duper.rule.elimDupLit] "Accessed nonsymmetry case" 48 | refs := refs.push (j, false) -- j is index of duplicate lit, false indicates the literal is not flipped 49 | | none => do 50 | match newLits.finIdxOf? (c.lits[i]!.symm) with 51 | | some j => do 52 | trace[duper.rule.elimDupLit] "Accessed symmetry case" 53 | refs := refs.push (j, true) -- j is index of duplicate lit, false indicates the literal is flipped 54 | | none => do 55 | refs := refs.push (newLits.size, false) 56 | newLits := newLits.push c.lits[i]! 57 | 58 | if newLits.size = c.lits.size then 59 | return none 60 | else 61 | let cp ← yieldClause (MClause.mk newLits) "eliminate duplicate literals" (some (mkElimDupLitProof refs)) 62 | return some #[cp] 63 | 64 | end Duper 65 | -------------------------------------------------------------------------------- /Duper/Rules/ElimResolvedLit.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | import Duper.Util.ProofReconstruction 3 | 4 | set_option linter.unusedVariables false 5 | 6 | namespace Duper 7 | open Std 8 | open RuleM 9 | open SimpResult 10 | open Lean 11 | open Meta 12 | 13 | initialize Lean.registerTraceClass `duper.rule.elimResolvedLit 14 | 15 | def mkElimResolvedLitProof (refs : List (Option Nat)) (premises : List Expr) (parents: List ProofParent) (transferExprs : Array Expr) 16 | (c : Clause) : MetaM Expr := do 17 | Meta.forallTelescope c.toForallExpr fun xs body => do 18 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 19 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 20 | let parentLits := parentsLits[0]! 21 | let appliedPremise := appliedPremises[0]! 22 | 23 | let mut proofCases : Array Expr := Array.mkEmpty parentLits.size 24 | for i in [:parentLits.size] do { 25 | let lit := parentLits[i]! 26 | if ((not lit.sign) && lit.lhs == lit.rhs) then 27 | -- lit has the form t ≠ t 28 | let proofCase ← Meta.withLocalDeclD `h lit.toExpr fun h => do 29 | let proofCase := mkApp2 (mkConst ``rfl [lit.lvl]) lit.ty lit.lhs 30 | let proofCase := mkApp h proofCase 31 | let proofCase := mkApp2 (mkConst ``False.elim [levelZero]) body proofCase 32 | Meta.mkLambdaFVars #[h] proofCase 33 | proofCases := proofCases.push proofCase 34 | else 35 | -- lit does not have the form t ≠ t, so refs[i] should have the value (some j) where parentLits[i] == c[j] 36 | match refs[i]! with 37 | | none => 38 | panic! "There is a bug in ElimResolvedLit.lean (The refs invariant is not satisfied)" 39 | | some j => 40 | let proofCase ← Meta.withLocalDeclD `h parentLits[i]!.toExpr fun h => do 41 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) j h 42 | proofCases := proofCases.push proofCase 43 | } 44 | let proof ← orCases (parentLits.map Lit.toExpr) proofCases 45 | Meta.mkLambdaFVars xs $ mkApp proof appliedPremise 46 | 47 | /-- Eliminate resolved literals (i.e. literals of the form t ≠ t) (Deletion of Resolved Literals: (DR)) -/ 48 | def elimResolvedLit : MSimpRule := fun c => do 49 | let c ← loadClause c 50 | /- 51 | Spec for newLits and refs: 52 | If c.lits[i] is resolved (i.e. of the form t ≠ t), then refs[i] = none 53 | If c.lits[i] isn't resolved, then refs[i] = some j where newLits[j] = c.lits[i] 54 | -/ 55 | let mut newLits : List Lit := [] 56 | let mut refs : List (Option Nat) := [] 57 | for lit in c.lits do { 58 | if ((not lit.sign) && lit.lhs == lit.rhs) then 59 | refs := none :: refs -- c.lits[i] is a resolved literal 60 | else 61 | refs := (some newLits.length) :: refs 62 | newLits := lit :: newLits -- Only add the literal if it does not have the form t ≠ t 63 | } 64 | -- To achieve the desired spec for newLits and refs, I must reverse them 65 | newLits := newLits.reverse 66 | refs := refs.reverse 67 | if (newLits.length = c.lits.size) then 68 | return none 69 | else 70 | trace[duper.rule.elimResolvedLit] "elimResolvedLit applied with the newLits: {newLits}" 71 | let yc ← yieldClause (MClause.mk newLits.toArray) "eliminate resolved literals" (some (mkElimResolvedLitProof refs)) 72 | return some #[yc] 73 | 74 | end Duper 75 | -------------------------------------------------------------------------------- /Duper/Rules/EqHoist.lean: -------------------------------------------------------------------------------- 1 | import Duper.MClause 2 | import Duper.RuleM 3 | import Duper.Simp 4 | import Duper.Util.ProofReconstruction 5 | 6 | set_option linter.unusedVariables false 7 | 8 | namespace Duper 9 | open Lean 10 | open Meta 11 | open RuleM 12 | open SimpResult 13 | 14 | initialize Lean.registerTraceClass `duper.rule.eqHoist 15 | 16 | theorem eq_hoist_proof (x y : α) (f : Prop → Prop) (h : f (x = y)) : f False ∨ x = y := by 17 | by_cases x_eq_y : x = y 18 | . exact Or.inr x_eq_y 19 | . rename ¬x = y => x_ne_y 20 | have x_eq_y_false := eq_false x_ne_y 21 | exact Or.inl $ x_eq_y_false ▸ h 22 | 23 | def mkEqHoistProof (pos : ClausePos) (premises : List Expr) 24 | (parents : List ProofParent) (transferExprs : Array Expr) (c : Clause) : MetaM Expr := 25 | Meta.forallTelescope c.toForallExpr fun xs body => do 26 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 27 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 28 | let parentLits := parentsLits[0]! 29 | let appliedPremise := appliedPremises[0]! 30 | 31 | let #[freshVar1, freshVar2] := transferExprs 32 | | throwError "mkEqHoistProof :: Wrong number of transferExprs" 33 | 34 | let mut caseProofs := Array.mkEmpty parentLits.size 35 | for i in [:parentLits.size] do 36 | let lit := parentLits[i]! 37 | let pr : Expr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 38 | if i == pos.lit then 39 | let substLitPos : LitPos := ⟨pos.side, pos.pos⟩ 40 | let abstrLit ← (lit.abstractAtPos! substLitPos) 41 | let abstrExp := abstrLit.toExpr 42 | let abstrLam := mkLambda `x BinderInfo.default (mkSort levelZero) abstrExp 43 | let lastTwoClausesProof ← Meta.mkAppM ``eq_hoist_proof #[freshVar1, freshVar2, abstrLam, h] 44 | Meta.mkLambdaFVars #[h] $ ← orSubclause (cLits.map Lit.toExpr) 2 lastTwoClausesProof 45 | else 46 | let idx := if i ≥ pos.lit then i - 1 else i 47 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) idx h 48 | caseProofs := caseProofs.push pr 49 | let r ← orCases (parentLits.map Lit.toExpr) caseProofs 50 | Meta.mkLambdaFVars xs $ mkApp r appliedPremise 51 | 52 | def eqHoistAtExpr (e : Expr) (pos : ClausePos) (given : Clause) (c : MClause) : RuleM (Array ClauseStream) := 53 | withoutModifyingMCtx do 54 | let lit := c.lits[pos.lit]! 55 | if e.getTopSymbol.isMVar' then -- Check condition 4 56 | -- If the head of e is a variable then it must be applied and the affected literal must be either 57 | -- e = True, e = False, or e = e' where e' is another variable headed term 58 | if not e.isApp then -- e is a non-applied variable and so we cannot apply eqHoist 59 | return #[] 60 | if pos.pos != #[] then 61 | return #[] -- e is not at the top level so the affected literal cannot have the form e = ... 62 | if not lit.sign then 63 | return #[] -- The affected literal is not positive and so it cannot have the form e = ... 64 | let otherSide := lit.getOtherSide pos.side 65 | if otherSide != (mkConst ``True) && otherSide != (mkConst ``False) && not otherSide.getTopSymbol.isMVar' then 66 | return #[] -- The other side is not True, False, or variable headed, so the affected literal cannot have the required form 67 | -- Check conditions 1 and 3 (condition 2 is guaranteed by construction) 68 | let eligibility ← eligibilityPreUnificationCheck c (alreadyReduced := true) pos.lit 69 | if eligibility == Eligibility.notEligible then 70 | return #[] 71 | -- Make freshVars and freshVarEquality 72 | let freshVar1 ← mkFreshExprMVar none 73 | let freshVarTy ← inferType freshVar1 74 | let freshVar2 ← mkFreshExprMVar freshVarTy 75 | let freshVarEquality ← mkAppM ``Eq #[freshVar1, freshVar2] 76 | -- Perform unification 77 | let ug ← unifierGenerator #[(e, freshVarEquality)] 78 | let loaded ← getLoadedClauses 79 | let yC := do 80 | setLoadedClauses loaded 81 | if not $ ← eligibilityPostUnificationCheck c (alreadyReduced := false) pos.lit eligibility (strict := lit.sign) then 82 | return none 83 | let eSide := lit.getSide pos.side 84 | let otherSide := lit.getOtherSide pos.side 85 | let cmp ← compare eSide otherSide false 86 | if cmp == Comparison.LessThan || cmp == Comparison.Equal then -- If eSide ≤ otherSide then e is not in an eligible position 87 | return none 88 | -- All side conditions have been met. Yield the appropriate clause 89 | let cErased := c.eraseLit pos.lit 90 | let some newLit ← lit.replaceAtPosUpdateType? ⟨pos.side, pos.pos⟩ (mkConst ``False) 91 | | return none -- If `lit` can't be safely changed at position `pos`, then don't apply `eqHoist` at `pos` 92 | let newClause := cErased.appendLits #[newLit, Lit.fromExpr freshVarEquality] 93 | trace[duper.rule.eqHoist] "Created {newClause.lits} from {c.lits}" 94 | let mkProof := mkEqHoistProof pos 95 | yieldClause newClause "eqHoist" mkProof (transferExprs := #[freshVar1, freshVar2]) 96 | return #[ClauseStream.mk ug given yC "eqHoist"] 97 | 98 | def eqHoist (given : Clause) (c : MClause) (cNum : Nat) : RuleM (Array ClauseStream) := do 99 | trace[duper.rule.eqHoist] "Running EqHoist on {c.lits}" 100 | let fold_fn := fun streams e pos => do 101 | let str ← eqHoistAtExpr e.consumeMData pos given c 102 | return streams.append str 103 | c.foldGreenM fold_fn #[] 104 | -------------------------------------------------------------------------------- /Duper/Rules/EqualityResolution.lean: -------------------------------------------------------------------------------- 1 | import Duper.RuleM 2 | import Duper.MClause 3 | import Duper.Util.ProofReconstruction 4 | 5 | set_option linter.unusedVariables false 6 | 7 | namespace Duper 8 | open RuleM 9 | open Lean 10 | open Meta 11 | 12 | initialize Lean.registerTraceClass `duper.rule.equalityResolution 13 | 14 | def mkEqualityResolutionProof (i : Nat) (premises : List Expr) (parents : List ProofParent) (transferExprs : Array Expr) 15 | (c : Clause) : MetaM Expr := 16 | Meta.forallTelescope c.toForallExpr fun xs body => do 17 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 18 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 19 | let parentLits := parentsLits[0]! 20 | let appliedPremise := appliedPremises[0]! 21 | 22 | let mut caseProofs := Array.mkEmpty parentLits.size 23 | for j in [:parentLits.size] do 24 | let lit := parentLits[j]! 25 | if j == i then 26 | -- lit has the form t ≠ t 27 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 28 | let pr := mkApp2 (mkConst ``rfl [lit.lvl]) lit.ty lit.lhs 29 | let pr := mkApp h pr 30 | let pr := mkApp2 (mkConst ``False.elim [levelZero]) body pr 31 | Meta.mkLambdaFVars #[h] pr 32 | caseProofs := caseProofs.push pr 33 | else 34 | -- need proof of `L_j → L_1 ∨ ... ∨ L_n` 35 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 36 | let idx := if j ≥ i then j - 1 else j 37 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) idx h 38 | -- caseProofs := caseProofs.push $ ← Lean.Meta.mkSorry (← Meta.mkArrow lit.toExpr body) (synthetic := true) 39 | caseProofs := caseProofs.push pr 40 | 41 | let r ← orCases (parentLits.map Lit.toExpr) caseProofs 42 | Meta.mkLambdaFVars xs $ mkApp r appliedPremise 43 | 44 | def equalityResolutionAtLit (given : Clause) (c : MClause) (i : Nat) : RuleM (Array ClauseStream) := 45 | withoutModifyingMCtx $ do 46 | let lit := c.lits[i]! 47 | let eligibility ← eligibilityPreUnificationCheck c (alreadyReduced := true) i 48 | if eligibility == Eligibility.notEligible then return #[] 49 | let loaded ← getLoadedClauses 50 | let ug ← unifierGenerator #[(lit.lhs, lit.rhs)] 51 | let yC := do 52 | setLoadedClauses loaded 53 | if not $ ← eligibilityPostUnificationCheck c (alreadyReduced := false) i eligibility then 54 | return none 55 | let c := c.eraseLit i 56 | some <$> yieldClause c "equality resolution" 57 | (mkProof := mkEqualityResolutionProof i) 58 | return #[ClauseStream.mk ug given yC "equality resolution"] 59 | 60 | def equalityResolution (given : Clause) (c : MClause) (cNum : Nat) : RuleM (Array ClauseStream) := do 61 | trace[duper.rule.equalityResolution] "EqRes inferences with {c.lits}" 62 | let mut streams := #[] 63 | for i in [:c.lits.size] do 64 | if c.lits[i]!.sign = false then 65 | streams := streams.append (← equalityResolutionAtLit given c i) 66 | return streams 67 | 68 | end Duper 69 | -------------------------------------------------------------------------------- /Duper/Rules/ExistsHoist.lean: -------------------------------------------------------------------------------- 1 | import Duper.MClause 2 | import Duper.RuleM 3 | import Duper.Simp 4 | import Duper.Util.ProofReconstruction 5 | 6 | set_option linter.unusedVariables false 7 | 8 | namespace Duper 9 | open Lean 10 | open Meta 11 | open RuleM 12 | open SimpResult 13 | 14 | initialize Lean.registerTraceClass `duper.rule.existsHoist 15 | 16 | theorem exists_hoist_proof {y : α → Prop} (x : α) (f : Prop → Prop) (h : f (∃ z : α, y z)) : f True ∨ y x = False := by 17 | by_cases z_hyp : ∃ z : α, y z 18 | . exact Or.inl (eq_true z_hyp ▸ h) 19 | . simp at z_hyp 20 | exact Or.inr (eq_false (z_hyp x)) 21 | 22 | def mkExistsHoistProof (pos : ClausePos) (premises : List Expr) (parents : List ProofParent) 23 | (transferExprs : Array Expr) (c : Clause) : MetaM Expr := 24 | Meta.forallTelescope c.toForallExpr fun xs body => do 25 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 26 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 27 | let parentLits := parentsLits[0]! 28 | let appliedPremise := appliedPremises[0]! 29 | 30 | let #[freshVar1] := transferExprs 31 | | throwError "mkExistsHoistProof :: Wrong number of transferExprs" 32 | 33 | let mut caseProofs := Array.mkEmpty parentLits.size 34 | for i in [:parentLits.size] do 35 | let lit := parentLits[i]! 36 | let pr : Expr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 37 | if i == pos.lit then 38 | let substLitPos : LitPos := ⟨pos.side, pos.pos⟩ 39 | let abstrLit ← (lit.abstractAtPos! substLitPos) 40 | let abstrExp := abstrLit.toExpr 41 | let abstrLam := mkLambda `x BinderInfo.default (mkSort levelZero) abstrExp 42 | let lastTwoLitsProof ← Meta.mkAppM ``exists_hoist_proof #[freshVar1, abstrLam, h] 43 | Meta.mkLambdaFVars #[h] $ ← orSubclause (cLits.map Lit.toExpr) 2 lastTwoLitsProof 44 | else 45 | let idx := if i ≥ pos.lit then i - 1 else i 46 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) idx h 47 | caseProofs := caseProofs.push pr 48 | let r ← orCases (parentLits.map Lit.toExpr) caseProofs 49 | Meta.mkLambdaFVars xs $ mkApp r appliedPremise 50 | 51 | def existsHoistAtExpr (e : Expr) (pos : ClausePos) (given : Clause) (c : MClause) : RuleM (Array ClauseStream) := 52 | withoutModifyingMCtx do 53 | let lit := c.lits[pos.lit]! 54 | if e.getTopSymbol.isMVar' then -- Check condition 4 55 | -- If the head of e is a variable then it must be applied and the affected literal must be either 56 | -- e = True, e = False, or e = e' where e' is another variable headed term 57 | if not e.isApp then -- e is a non-applied variable and so we cannot apply neHoist 58 | return #[] 59 | if pos.pos != #[] then 60 | return #[] -- e is not at the top level so the affected literal cannot have the form e = ... 61 | if not lit.sign then 62 | return #[] -- The affected literal is not positive and so it cannot have the form e = ... 63 | let otherSide := lit.getOtherSide pos.side 64 | if otherSide != (mkConst ``True) && otherSide != (mkConst ``False) && not otherSide.getTopSymbol.isMVar' then 65 | return #[] -- The other side is not True, False, or variable headed, so the affected literal cannot have the required form 66 | -- Check conditions 1 and 3 (condition 2 is guaranteed by construction) 67 | let eligibility ← eligibilityPreUnificationCheck c (alreadyReduced := true) pos.lit 68 | if eligibility == Eligibility.notEligible then 69 | return #[] 70 | -- Make freshVars, freshVarExistsExpr and newLitLhs 71 | let freshVar1 ← mkFreshExprMVar none 72 | let freshVar1Ty ← inferType freshVar1 73 | let freshVar2Ty := Expr.forallE .anonymous freshVar1Ty (mkSort levelZero) BinderInfo.default -- freshVar1Ty → Prop 74 | let freshVar2 ← mkFreshExprMVar freshVar2Ty 75 | let freshVarExistsExpr ← mkAppM ``Exists #[freshVar2] 76 | let newLitLhs := .app freshVar2 freshVar1 77 | -- Perform unification 78 | let ug ← unifierGenerator #[(e, freshVarExistsExpr)] 79 | let loaded ← getLoadedClauses 80 | let yC := do 81 | setLoadedClauses loaded 82 | if not $ ← eligibilityPostUnificationCheck c (alreadyReduced := false) pos.lit eligibility (strict := lit.sign) then 83 | return none 84 | let eSide := lit.getSide pos.side 85 | let otherSide := lit.getOtherSide pos.side 86 | let cmp ← compare eSide otherSide false 87 | if cmp == Comparison.LessThan || cmp == Comparison.Equal then -- If eSide ≤ otherSide then e is not in an eligible position 88 | return none 89 | -- All side conditions have been met. Yield the appropriate clause 90 | let cErased := c.eraseLit pos.lit 91 | let some replacedLit ← lit.replaceAtPosUpdateType? ⟨pos.side, pos.pos⟩ (mkConst ``True) 92 | | return none 93 | let newClause := cErased.appendLits #[replacedLit, Lit.fromSingleExpr newLitLhs (sign := false)] 94 | trace[duper.rule.existsHoist] "Created {newClause.lits} from {c.lits}" 95 | yieldClause newClause "existsHoist" (some (mkExistsHoistProof pos)) (transferExprs := #[freshVar1]) 96 | return #[ClauseStream.mk ug given yC "existsHoist"] 97 | 98 | def existsHoist (given : Clause) (c : MClause) (cNum : Nat) : RuleM (Array ClauseStream) := do 99 | trace[duper.rule.existsHoist] "Running ExistsHoist on {c.lits}" 100 | let fold_fn := fun streams e pos => do 101 | let str ← existsHoistAtExpr e.consumeMData pos given c 102 | return streams.append str 103 | c.foldGreenM fold_fn #[] 104 | -------------------------------------------------------------------------------- /Duper/Rules/FalseElim.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | import Duper.Util.ProofReconstruction 3 | import Lean.Meta.Basic 4 | 5 | set_option linter.unusedVariables false 6 | 7 | namespace Duper 8 | open RuleM 9 | open SimpResult 10 | open Lean 11 | open Meta 12 | 13 | initialize registerTraceClass `duper.rule.falseElim 14 | 15 | /-- Determines whether a literal has can be unified with the form `False = True` or `True = False`. If it can 16 | be, then the substitution necessary to achieve that match is applied (if the match is unsuccessful, then 17 | the MCtx remains unchanged) -/ 18 | def isFalseLiteral (lit : Lit) : MetaM Bool := do 19 | if ← Meta.fastUnify #[(lit.lhs, mkConst ``False), (lit.rhs, mkConst ``True)] then return true 20 | else if ← Meta.fastUnify #[(lit.lhs, mkConst ``True), (lit.rhs, mkConst ``False)] then return true 21 | else return false 22 | 23 | theorem false_ne_true (h : False = True) : False := by rw [h]; exact ⟨⟩ 24 | theorem true_ne_false (h : True = False) : False := by rw [← h]; exact ⟨⟩ 25 | 26 | def mkFalseElimProof (i : Nat) (premises : List Expr) (parents : List ProofParent) (transferExprs : Array Expr) 27 | (c : Clause) : MetaM Expr := 28 | Meta.forallTelescope c.toForallExpr fun xs body => do 29 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 30 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 31 | let parentLits := parentsLits[0]! 32 | let appliedPremise := appliedPremises[0]! 33 | 34 | let mut caseProofs := Array.mkEmpty parentLits.size 35 | for j in [:parentLits.size] do 36 | let lit := parentLits[j]! 37 | if i == j then 38 | if lit.lhs == mkConst ``True then -- lit unified with `True = False` 39 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 40 | let proofCase := mkApp (mkConst ``true_ne_false) h 41 | let proofCase := mkApp2 (mkConst ``False.elim [levelZero]) body proofCase 42 | Meta.mkLambdaFVars #[h] proofCase 43 | caseProofs := caseProofs.push pr 44 | else -- lit unified with `False = True` 45 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 46 | let proofCase := mkApp (mkConst ``false_ne_true) h 47 | let proofCase := mkApp2 (mkConst ``False.elim [levelZero]) body proofCase 48 | Meta.mkLambdaFVars #[h] proofCase 49 | caseProofs := caseProofs.push pr 50 | else 51 | -- need proof of `L_j → L_1 ∨ ... ∨ L_n` 52 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 53 | let idx := if j ≥ i then j - 1 else j 54 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) idx h 55 | caseProofs := caseProofs.push pr 56 | 57 | let r ← orCases (parentLits.map Lit.toExpr) caseProofs 58 | Meta.mkLambdaFVars xs $ mkApp r appliedPremise 59 | 60 | def falseElimAtLit (given : Clause) (c : MClause) (i : Nat) : RuleM (Array ClauseStream) := 61 | withoutModifyingMCtx $ do 62 | let lit := c.lits[i]! 63 | let eligibility ← eligibilityPreUnificationCheck c (alreadyReduced := true) i 64 | if eligibility == Eligibility.notEligible then return #[] 65 | let loaded ← getLoadedClauses 66 | let ug ← DUnif.UnifierGenerator.fromMetaMProcedure (isFalseLiteral lit) 67 | let yC := do 68 | setLoadedClauses loaded 69 | if (not $ ← eligibilityPostUnificationCheck c (alreadyReduced := false) i eligibility (strict := true)) then return none 70 | let c := c.eraseLit i 71 | trace[duper.rule.falseElim] "Successfully yielded {c.lits} by removing literal {i}" 72 | yieldClause c "falseElim" $ some (mkFalseElimProof i) 73 | return #[ClauseStream.mk ug given yC "falseElim"] 74 | 75 | /-- If there is a substitution σ and literal l in c such that σ(l) equates `True` and `False`, then 76 | falseElim yields the clause obtained by removing l from c and applying sigma to the rest of c -/ 77 | def falseElim (given : Clause) (c : MClause) (cNum : Nat) : RuleM (Array ClauseStream) := do 78 | trace[duper.rule.falseElim] "Running FalseElim on {c.lits}" 79 | let mut streams := #[] 80 | for i in [:c.lits.size] do 81 | if c.lits[i]!.sign then 82 | let str ← falseElimAtLit given c i 83 | streams := streams.append str 84 | return streams 85 | 86 | end Duper 87 | -------------------------------------------------------------------------------- /Duper/Rules/IdentBoolFalseElim.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | import Duper.Util.ProofReconstruction 3 | 4 | set_option linter.unusedVariables false 5 | 6 | namespace Duper 7 | open RuleM 8 | open SimpResult 9 | open Lean 10 | open Meta 11 | 12 | /-- Determines whether a literal has exactly the form `false = true` or `true = false`-/ 13 | def isFalseBoolLiteral (lit : Lit) : MetaM Bool := do 14 | if lit.ty.consumeMData == (mkConst ``Bool) then 15 | return lit.sign && 16 | ((lit.lhs == mkConst ``true && lit.rhs == mkConst ``false) || 17 | (lit.lhs == mkConst ``false && lit.rhs == mkConst ``true)) 18 | else return false 19 | 20 | theorem bool_false_ne_true (h : false = true) : False := ne_false_of_eq_true h (by rfl) 21 | 22 | theorem bool_true_ne_false (h : true = false) : False := ne_true_of_eq_false h (by rfl) 23 | 24 | def mkIdentBoolFalseElimProof (refs : List (Option Nat)) (premises : List Expr) (parents: List ProofParent) 25 | (transferExprs : Array Expr) (c : Clause) : MetaM Expr := 26 | Meta.forallTelescope c.toForallExpr fun xs body => do 27 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 28 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 29 | let parentLits := parentsLits[0]! 30 | let appliedPremise := appliedPremises[0]! 31 | 32 | let mut proofCases : Array Expr := Array.mkEmpty parentLits.size 33 | for i in [:parentLits.size] do 34 | let lit := parentLits[i]! 35 | if (← isFalseBoolLiteral lit) then -- lit has the form `false = true` or `true = false` 36 | let proofCase ← Meta.withLocalDeclD `h lit.toExpr fun h => do 37 | if (lit.lhs == mkConst ``false) then 38 | let proofCase := mkApp (mkConst ``bool_false_ne_true) h 39 | let proofCase := mkApp2 (mkConst ``False.elim [levelZero]) body proofCase 40 | Meta.mkLambdaFVars #[h] proofCase 41 | else if(lit.lhs == mkConst ``true) then 42 | let proofCase := mkApp (mkConst ``bool_true_ne_false) h 43 | let proofCase := mkApp2 (mkConst ``False.elim [levelZero]) body proofCase 44 | Meta.mkLambdaFVars #[h] proofCase 45 | else 46 | throwError "mkIdentBoolFalseElimProof failed to match {lit.lhs} to an expected expression" 47 | proofCases := proofCases.push proofCase 48 | else -- refs[i] should have the value (some j) where parentLits[i] == c[j] 49 | match refs[i]! with 50 | | none => throwError "Refs invariant is not satisfied in identBoolFalseElim" 51 | | some j => 52 | let proofCase ← Meta.withLocalDeclD `h parentLits[i]!.toExpr fun h => do 53 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) j h 54 | proofCases := proofCases.push proofCase 55 | let proof ← orCases (parentLits.map Lit.toExpr) proofCases 56 | Meta.mkLambdaFVars xs $ mkApp proof appliedPremise 57 | 58 | /-- Eliminate literals that are exactly of the form `false = true` or `true = false`. 59 | This is a special case of the boolFalseElim inference rule in which σ is the identity. 60 | This rule is included as a means of giving Bools special attention. -/ 61 | def identBoolFalseElim : MSimpRule := fun c => do 62 | let c ← loadClause c 63 | /- 64 | Spec for newLits and refs 65 | If c.lits[i] is `false = true` or `true = false`, then refs[i] = none 66 | If c.lits[i] isn't `false = true` or `true = false`,then refs[i] = some j where newLits[j] = c.lits[i] 67 | -/ 68 | let mut newLits : List Lit := [] 69 | let mut refs : List (Option Nat) := [] 70 | for lit in c.lits do 71 | if (← isFalseBoolLiteral lit) then 72 | refs := none :: refs 73 | else 74 | refs := (some newLits.length) :: refs 75 | newLits := lit :: newLits 76 | -- To achieve the desired spec for newLits and refs, I must reverse them 77 | newLits := newLits.reverse 78 | refs := refs.reverse 79 | if (newLits.length = c.lits.size) then 80 | return none 81 | else 82 | let resultClause ← yieldClause (MClause.mk newLits.toArray) "identity boolean false elimination" 83 | (some (mkIdentBoolFalseElimProof refs)) 84 | return some #[resultClause] 85 | 86 | end Duper 87 | -------------------------------------------------------------------------------- /Duper/Rules/IdentBoolHoist.lean: -------------------------------------------------------------------------------- 1 | import Duper.MClause 2 | import Duper.RuleM 3 | import Duper.Simp 4 | import Duper.Util.ProofReconstruction 5 | 6 | set_option linter.unusedVariables false 7 | 8 | -- LoobHoist BoolHoist 9 | -- C 10 | -- ----------------------------- 11 | -- C<⊤> ∨ u = ⊥ C<⊥> ∨ u = ⊤ 12 | -- 13 | -- where u is of type Prop, but neither ⊥ nor ⊤, and u is not at the top level of a positive literal. 14 | 15 | namespace Duper 16 | open Lean 17 | open RuleM 18 | open Meta 19 | open SimpResult 20 | 21 | initialize Lean.registerTraceClass `duper.rule.identBoolHoist 22 | 23 | theorem loob_hoist_proof (f : Prop → Prop) (e : Prop) (H : f e) : f True ∨ e = False := 24 | @Classical.byCases e _ 25 | (fun p => have h : e = True := by simp [p]; 26 | Or.inl (h ▸ H)) 27 | (fun np => by simp [np]) 28 | 29 | theorem bool_hoist_proof (f : Prop → Prop) (e : Prop) (H : f e) : f False ∨ e = True := 30 | @Classical.byCases e _ 31 | (fun p => by simp [p]) 32 | (fun np => have h : e = False := by simp [np]; 33 | Or.inl (h ▸ H)) 34 | 35 | def mkBoolHoistProof (pos : ClausePos) (sgn : Bool) (premises : List Expr) 36 | (parents : List ProofParent) (transferExprs : Array Expr) (c : Clause) : MetaM Expr := 37 | Meta.forallTelescope c.toForallExpr fun xs body => do 38 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 39 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 40 | let parentLits := parentsLits[0]! 41 | let appliedPremise := appliedPremises[0]! 42 | 43 | let i := pos.lit 44 | let ⟨s, p⟩ := pos.toLitPos 45 | 46 | let mut caseProofs := Array.mkEmpty parentLits.size 47 | for j in [:parentLits.size] do 48 | let lit := parentLits[j]! 49 | if j == i then 50 | let lp : LitPos := ⟨s, p⟩ 51 | let f ← Meta.withLocalDeclD `h (.sort .zero) fun h => do 52 | let lit' ← lit.replaceAtPos! lp h 53 | let f := lit'.toExpr 54 | Meta.mkLambdaFVars #[h] f 55 | let e ← lit.getAtPos! lp 56 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 57 | let mut pr := h 58 | if sgn then 59 | pr ← Meta.mkAppM ``loob_hoist_proof #[f, e, h] 60 | else 61 | pr ← Meta.mkAppM ``bool_hoist_proof #[f, e, h] 62 | Meta.mkLambdaFVars #[h] $ ← orSubclause (cLits.map Lit.toExpr) 2 pr 63 | caseProofs := caseProofs.push pr 64 | else 65 | -- need proof of `L_j → L_1 ∨ ... ∨ L_n` 66 | let pr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 67 | let idx := if j ≥ i then j - 1 else j 68 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) idx h 69 | caseProofs := caseProofs.push pr 70 | 71 | let r ← orCases (parentLits.map Lit.toExpr) caseProofs 72 | Meta.mkLambdaFVars xs $ mkApp r appliedPremise 73 | 74 | def identBoolHoistAtExpr (e : Expr) (pos : ClausePos) (c : MClause) : RuleM (Option (Array (Clause × Proof))) := 75 | withoutModifyingMCtx do 76 | let ty ← inferType e 77 | if ty == .sort .zero then 78 | let l := pos.lit 79 | let ⟨s, p⟩ := pos.toLitPos 80 | trace[duper.rule.identBoolHoist] m!"Inspecting position {pos} in clause {c.lits.map Lit.toExpr}" 81 | let is_true := e == (mkConst ``True) 82 | let is_false := e == (mkConst ``False) 83 | let is_top_positive := p.size == 0 ∧ c.lits[l]!.sign 84 | if is_true ∨ is_false ∨ is_top_positive then 85 | return none 86 | else 87 | trace[duper.rule.identBoolHoist] m!"BoolHoist at literal {l}, side {s}, position {p} in clause {c.lits.map Lit.toExpr}" 88 | let litl := c.lits[l]! 89 | let some litl1 ← litl.replaceAtPosUpdateType? ⟨s, p⟩ (mkConst ``True) 90 | | return none -- If `litl` can't be safely changed at position `p`, then don't apply `identBoolHoist` at `p` 91 | let some litl2 ← litl.replaceAtPosUpdateType? ⟨s, p⟩ (mkConst ``False) 92 | | return none -- If `litl` can't be safely changed at position `p`, then don't apply `identBoolHoist` at `p` 93 | let c_erased := c.eraseLit l 94 | let nc := c_erased.appendLits #[litl1, Lit.fromSingleExpr e false] 95 | trace[duper.rule.identBoolHoist] s!"New Clause: {nc.lits.map Lit.toExpr}" 96 | let cp1 ← yieldClause nc "identity loobHoist" (some (mkBoolHoistProof pos true)) 97 | let nc := c_erased.appendLits #[litl2, Lit.fromSingleExpr e true] 98 | trace[duper.rule.identBoolHoist] s!"New Clause: {nc.lits.map Lit.toExpr}" 99 | let cp2 ← yieldClause nc "identity boolHoist" (some (mkBoolHoistProof pos false)) 100 | return some #[cp1, cp2] 101 | else 102 | return none 103 | 104 | def identBoolHoist : MSimpRule := fun c => do 105 | let c ← loadClause c 106 | let fold_fn := fun acc e pos => do 107 | match acc with 108 | | some res => return some res 109 | | none => identBoolHoistAtExpr e pos c 110 | c.foldGreenM fold_fn none 111 | -------------------------------------------------------------------------------- /Duper/Rules/IdentPropFalseElim.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | import Duper.Util.ProofReconstruction 3 | import Lean.Meta.Basic 4 | 5 | set_option linter.unusedVariables false 6 | 7 | namespace Duper 8 | open RuleM 9 | open SimpResult 10 | open Lean 11 | open Meta 12 | 13 | initialize Lean.registerTraceClass `duper.rule.identPropFalseElim 14 | 15 | /-- Determines whether a literal has exactly the form `False = True` or `True = False`-/ 16 | def isFalsePropLiteral (lit : Lit) : MetaM Bool := do 17 | match lit.ty with 18 | | Expr.sort lvl => 19 | if Level.isEquiv (← Lean.instantiateLevelMVars lvl) levelZero then 20 | return lit.sign && 21 | ((lit.lhs == mkConst ``True && lit.rhs == mkConst ``False) || 22 | (lit.lhs == mkConst ``False && lit.rhs == mkConst ``True)) 23 | else return false 24 | | _ => return false 25 | 26 | theorem prop_false_ne_true (h : False = True) : False := by rw [h]; exact ⟨⟩ 27 | 28 | theorem prop_true_ne_false (h : True = False) : False := by rw [← h]; exact ⟨⟩ 29 | 30 | def mkIdentPropFalseElimProof (refs : List (Option Nat)) (premises : List Expr) (parents : List ProofParent) (transferExprs : Array Expr) 31 | (c : Clause) : MetaM Expr := 32 | Meta.forallTelescope c.toForallExpr fun xs body => do 33 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 34 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 35 | let parentLits := parentsLits[0]! 36 | let appliedPremise := appliedPremises[0]! 37 | 38 | let mut proofCases : Array Expr := Array.mkEmpty parentLits.size 39 | for i in [:parentLits.size] do 40 | let lit := parentLits[i]! 41 | if (← isFalsePropLiteral lit) then -- lit has the form `False = True` or `True = False` 42 | let proofCase ← Meta.withLocalDeclD `h lit.toExpr fun h => do 43 | if (lit.lhs == mkConst ``False) then 44 | let proofCase := mkApp (mkConst ``prop_false_ne_true) h 45 | let proofCase := mkApp2 (mkConst ``False.elim [levelZero]) body proofCase 46 | Meta.mkLambdaFVars #[h] proofCase 47 | else if(lit.lhs == mkConst ``True) then 48 | let proofCase := mkApp (mkConst ``prop_true_ne_false) h 49 | let proofCase := mkApp2 (mkConst ``False.elim [levelZero]) body proofCase 50 | Meta.mkLambdaFVars #[h] proofCase 51 | else 52 | throwError "mkIdentPropFalseElimProof failed to match {lit.lhs} to an expected expression" 53 | proofCases := proofCases.push proofCase 54 | else -- refs[i] should have the value (some j) where parentLits[i] == c[j] 55 | match refs[i]! with 56 | | none => throwError "Refs invariant is not satisfied in identPropFalseElim" 57 | | some j => 58 | let proofCase ← Meta.withLocalDeclD `h parentLits[i]!.toExpr fun h => do 59 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) j h 60 | proofCases := proofCases.push proofCase 61 | let proof ← orCases (parentLits.map Lit.toExpr) proofCases 62 | Meta.mkLambdaFVars xs $ mkApp proof appliedPremise 63 | 64 | /-- Eliminate literals that are exactly of the form `False = True` or `True = False`. 65 | This is a special case of the propFalseElim inference rule in which σ is the identity. -/ 66 | def identPropFalseElim : MSimpRule := fun c => do 67 | let c ← loadClause c 68 | /- 69 | Spec for newLits and refs: 70 | If c.lits[i] is `False = True` or `True = False`, then refs[i] = none 71 | If c.lits[i] isn't `False = True` or `True = False`,then refs[i] = some j where newLits[j] = c.lits[i] 72 | -/ 73 | let mut newLits : List Lit := [] 74 | let mut refs : List (Option Nat) := [] 75 | for lit in c.lits do 76 | if (← isFalsePropLiteral lit) then 77 | refs := none :: refs 78 | else 79 | refs := (some newLits.length) :: refs 80 | newLits := lit :: newLits 81 | -- To achieve the desired spec for newLits and refs, I must reverse them 82 | newLits := newLits.reverse 83 | refs := refs.reverse 84 | if (newLits.length = c.lits.size) then 85 | trace[duper.rule.identPropFalseElim] "Returning Unapplicable on {c.lits}" 86 | return none 87 | else 88 | trace[duper.rule.identPropFalseElim] "Succeeded on {c.lits}, yielding {newLits}" 89 | let resultClause ← yieldClause (MClause.mk newLits.toArray) "identity prop false elimination" 90 | (some (mkIdentPropFalseElimProof refs)) 91 | return some #[resultClause] 92 | 93 | end Duper 94 | -------------------------------------------------------------------------------- /Duper/Rules/NeHoist.lean: -------------------------------------------------------------------------------- 1 | import Duper.MClause 2 | import Duper.RuleM 3 | import Duper.Simp 4 | import Duper.Util.ProofReconstruction 5 | 6 | set_option linter.unusedVariables false 7 | 8 | namespace Duper 9 | open Lean 10 | open Meta 11 | open RuleM 12 | open SimpResult 13 | 14 | initialize Lean.registerTraceClass `duper.rule.neHoist 15 | 16 | theorem ne_hoist_proof (x y : α) (f : Prop → Prop) (h : f (x ≠ y)) : f True ∨ x = y := by 17 | by_cases x_eq_y : x = y 18 | . exact Or.inr x_eq_y 19 | . rename ¬x = y => x_ne_y 20 | have x_ne_y_true := eq_true x_ne_y 21 | exact Or.inl $ x_ne_y_true ▸ h 22 | 23 | def mkNeHoistProof (pos : ClausePos) (premises : List Expr) 24 | (parents : List ProofParent) (transferExprs : Array Expr) (c : Clause) : MetaM Expr := 25 | Meta.forallTelescope c.toForallExpr fun xs body => do 26 | let cLits := c.lits.map (fun l => l.map (fun e => e.instantiateRev xs)) 27 | let (parentsLits, appliedPremises, transferExprs) ← instantiatePremises parents premises xs transferExprs 28 | let parentLits := parentsLits[0]! 29 | let appliedPremise := appliedPremises[0]! 30 | 31 | let #[freshVar1, freshVar2] := transferExprs 32 | | throwError "mkNeHoistProof :: Wrong number of transferExprs" 33 | 34 | let mut caseProofs := Array.mkEmpty parentLits.size 35 | for i in [:parentLits.size] do 36 | let lit := parentLits[i]! 37 | let pr : Expr ← Meta.withLocalDeclD `h lit.toExpr fun h => do 38 | if i == pos.lit then 39 | let substLitPos : LitPos := ⟨pos.side, pos.pos⟩ 40 | let abstrLit ← (lit.abstractAtPos! substLitPos) 41 | let abstrExp := abstrLit.toExpr 42 | let abstrLam := mkLambda `x BinderInfo.default (mkSort levelZero) abstrExp 43 | let lastTwoClausesProof ← Meta.mkAppM ``ne_hoist_proof #[freshVar1, freshVar2, abstrLam, h] 44 | Meta.mkLambdaFVars #[h] $ ← orSubclause (cLits.map Lit.toExpr) 2 lastTwoClausesProof 45 | else 46 | let idx := if i ≥ pos.lit then i - 1 else i 47 | Meta.mkLambdaFVars #[h] $ ← orIntro (cLits.map Lit.toExpr) idx h 48 | caseProofs := caseProofs.push pr 49 | let r ← orCases (parentLits.map Lit.toExpr) caseProofs 50 | Meta.mkLambdaFVars xs $ mkApp r appliedPremise 51 | 52 | def neHoistAtExpr (e : Expr) (pos : ClausePos) (given : Clause) (c : MClause) : RuleM (Array ClauseStream) := 53 | withoutModifyingMCtx do 54 | let lit := c.lits[pos.lit]! 55 | if e.getTopSymbol.isMVar' then -- Check condition 4 56 | -- If the head of e is a variable then it must be applied and the affected literal must be either 57 | -- e = True, e = False, or e = e' where e' is another variable headed term 58 | if not e.isApp then -- e is a non-applied variable and so we cannot apply neHoist 59 | return #[] 60 | if pos.pos != #[] then 61 | return #[] -- e is not at the top level so the affected literal cannot have the form e = ... 62 | if not lit.sign then 63 | return #[] -- The affected literal is not positive and so it cannot have the form e = ... 64 | let otherSide := lit.getOtherSide pos.side 65 | if otherSide != (mkConst ``True) && otherSide != (mkConst ``False) && not otherSide.getTopSymbol.isMVar' then 66 | return #[] -- The other side is not True, False, or variable headed, so the affected literal cannot have the required form 67 | -- Check conditions 1 and 3 (condition 2 is guaranteed by construction) 68 | let eligibility ← eligibilityPreUnificationCheck c (alreadyReduced := true) pos.lit 69 | if eligibility == Eligibility.notEligible then 70 | return #[] 71 | -- Make freshVars, freshVarInequality, and freshVarEquality 72 | let freshVar1 ← mkFreshExprMVar none 73 | let freshVarTy ← inferType freshVar1 74 | let freshVar2 ← mkFreshExprMVar freshVarTy 75 | let freshVarInequality ← mkAppM ``Ne #[freshVar1, freshVar2] 76 | let freshVarEquality ← mkAppM ``Eq #[freshVar1, freshVar2] 77 | -- Perform unification 78 | let ug ← unifierGenerator #[(e, freshVarInequality)] 79 | let loaded ← getLoadedClauses 80 | let yC := do 81 | setLoadedClauses loaded 82 | if not $ ← eligibilityPostUnificationCheck c (alreadyReduced := false) pos.lit eligibility (strict := lit.sign) then 83 | return none 84 | let eSide := lit.getSide pos.side 85 | let otherSide := lit.getOtherSide pos.side 86 | let cmp ← compare eSide otherSide false 87 | if cmp == Comparison.LessThan || cmp == Comparison.Equal then -- If eSide ≤ otherSide then e is not in an eligible position 88 | return none 89 | -- All side conditions have been met. Yield the appropriate clause 90 | let cErased := c.eraseLit pos.lit 91 | let some newLit ← lit.replaceAtPosUpdateType? ⟨pos.side, pos.pos⟩ (mkConst ``True) 92 | | return none -- If `lit` can't be safely changed at position `pos`, then don't apply `neHoist` at `pos` 93 | let newClause := cErased.appendLits #[newLit, Lit.fromExpr freshVarEquality] 94 | trace[duper.rule.neHoist] "Created {newClause.lits} from {c.lits}" 95 | let mkProof := mkNeHoistProof pos 96 | yieldClause newClause "neHoist" mkProof (transferExprs := #[freshVar1, freshVar2]) 97 | return #[ClauseStream.mk ug given yC "neHoist"] 98 | 99 | def neHoist (given : Clause) (c : MClause) (cNum : Nat) : RuleM (Array ClauseStream) := do 100 | trace[duper.rule.neHoist] "Running NeHoist on {c.lits}" 101 | let fold_fn := fun streams e pos => do 102 | let str ← neHoistAtExpr e.consumeMData pos given c 103 | return streams.append str 104 | c.foldGreenM fold_fn #[] 105 | -------------------------------------------------------------------------------- /Duper/Rules/SyntacticTautologyDeletion1.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | 3 | namespace Duper 4 | open RuleM 5 | open SimpResult 6 | 7 | -- TODO: Do this on Clause instead of MClause? 8 | -- This implements the rule syntactic tautology deletion 1 (TD1) 9 | def syntacticTautologyDeletion1 : MSimpRule := fun c => do 10 | for lit in c.lits do 11 | if lit.sign ∧ lit.lhs == lit.rhs then 12 | return some #[] 13 | return none 14 | 15 | end Duper -------------------------------------------------------------------------------- /Duper/Rules/SyntacticTautologyDeletion2.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | 3 | namespace Duper 4 | open Lean 5 | open RuleM 6 | open SimpResult 7 | 8 | initialize Lean.registerTraceClass `duper.rule.syntacticTautologyDeletion 9 | 10 | /-- This implements the rule syntactic tautology deletion 2 (TD2) -/ 11 | def syntacticTautologyDeletion2 : MSimpRule := fun c => do 12 | let mut eq_pairs : Std.HashSet (Lean.Expr × Lean.Expr) := Std.HashSet.emptyWithCapacity -- A HashSet that stores pairs of Lean expressions that are equated in the clause 13 | let mut ne_pairs : Std.HashSet (Lean.Expr × Lean.Expr) := Std.HashSet.emptyWithCapacity -- A HashSet that stores pairs of Lean expressions that are stated to not be equal in the clause 14 | for lit in c.lits do 15 | let lhs := lit.lhs 16 | let rhs := lit.rhs 17 | if lit.sign then { 18 | if(ne_pairs.contains (lhs, rhs) || ne_pairs.contains (rhs, lhs)) then 19 | trace[duper.rule.syntacticTautologyDeletion] "syntacticTautologyDeletion2 returning Removed due to literals {lhs} and {rhs} from the clause {c.lits}" 20 | return some #[] -- The literal lhs = rhs and the literal lhs ≠ rhs are both in the clause, so the clause can be removed 21 | else 22 | eq_pairs := eq_pairs.insert (lhs, rhs) 23 | } 24 | else { 25 | if(eq_pairs.contains (lhs, rhs) || eq_pairs.contains (rhs, lhs)) then 26 | trace[duper.rule.syntacticTautologyDeletion] "syntacticTautologyDeletion2 returning Removed due to literals {lhs} and {rhs} from the clause {c.lits}" 27 | return some #[] -- The literal lhs ≠ rhs and the literal lhs = rhs are both in the clause, so the clause can be removed 28 | else 29 | ne_pairs := ne_pairs.insert (lhs, rhs) 30 | } 31 | return none 32 | 33 | end Duper 34 | -------------------------------------------------------------------------------- /Duper/Rules/SyntacticTautologyDeletion3.lean: -------------------------------------------------------------------------------- 1 | import Duper.Simp 2 | 3 | namespace Duper 4 | open Lean 5 | open Std 6 | open RuleM 7 | open SimpResult 8 | 9 | /-- Syntactic tautology deletion 3 doesn't refer to a specific rule in the literature, but is a response to the observation that 10 | duper often gets cluttered with clauses of the form "x = True ∨ x = False". Neither syntacticTautologyDeletion1 nor 11 | syntacticTautologyDeletion2 remove clauses of this form, so that is what syntacticTautologyDeletion3 targets. -/ 12 | def syntacticTautologyDeletion3 : MSimpRule := fun c => do 13 | let mut eqTrueSet : Std.HashSet Lean.Expr := Std.HashSet.emptyWithCapacity -- Stores Lean expressions equated to True in c 14 | let mut eqFalseSet : Std.HashSet Lean.Expr := Std.HashSet.emptyWithCapacity -- Stores Lean expressions equated to False in c 15 | for lit in c.lits do 16 | if lit.sign then 17 | if lit.rhs == mkConst ``True then 18 | if eqFalseSet.contains lit.lhs then return some #[] 19 | else eqTrueSet := eqTrueSet.insert lit.lhs 20 | else if lit.rhs == mkConst ``False then 21 | if eqTrueSet.contains lit.lhs then return some #[] 22 | else eqFalseSet := eqFalseSet.insert lit.lhs 23 | else if lit.lhs == mkConst ``True then 24 | if eqFalseSet.contains lit.rhs then return some #[] 25 | else eqTrueSet := eqTrueSet.insert lit.rhs 26 | else if lit.lhs == mkConst ``False then 27 | if eqTrueSet.contains lit.rhs then return some #[] 28 | else eqFalseSet := eqFalseSet.insert lit.rhs 29 | else -- We can view "x ≠ True" as the same as "x = False" 30 | if lit.rhs == mkConst ``True then 31 | if eqTrueSet.contains lit.lhs then return some #[] 32 | else eqFalseSet := eqFalseSet.insert lit.lhs 33 | else if lit.rhs == mkConst ``False then 34 | if eqFalseSet.contains lit.lhs then return some #[] 35 | else eqTrueSet := eqTrueSet.insert lit.lhs 36 | else if lit.lhs == mkConst ``True then 37 | if eqTrueSet.contains lit.rhs then return some #[] 38 | else eqFalseSet := eqFalseSet.insert lit.rhs 39 | else if lit.lhs == mkConst ``False then 40 | if eqFalseSet.contains lit.rhs then return some #[] 41 | else eqTrueSet := eqTrueSet.insert lit.rhs 42 | return none 43 | 44 | end Duper 45 | -------------------------------------------------------------------------------- /Duper/TPTP.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import Duper.TPTPParser.MacroDecl 3 | 4 | open Lean 5 | open Lean.Parser 6 | open TSyntax.Compat 7 | open Lean.Elab.Command 8 | 9 | namespace TPTP 10 | 11 | register_option maxTPTPProblemLines : Nat := { 12 | defValue := 10000 13 | descr := "Line number limit (with comments stripped) for TPTP problems" 14 | } 15 | 16 | def getMaxTPTPProblemLines (opts : Options) : Nat := 17 | maxTPTPProblemLines.get opts 18 | 19 | def checkMaxTPTPProblemLines (lines : Nat) : CommandElabM Unit := do 20 | let opts ← getOptions 21 | let max := getMaxTPTPProblemLines opts 22 | if max < lines then 23 | let msg := s!"Number of lines {lines} in TPTP problem exceeded line number limit {max}" 24 | throw <| Exception.error (← getRef) (MessageData.ofFormat (Std.Format.text msg)) 25 | 26 | partial def parseTPTPInput (s : String) : CommandElabM Syntax := do 27 | match runParserCategory (← getEnv) `TPTP_input s with 28 | | Except.error e => throwError e 29 | | Except.ok r => return r 30 | 31 | def sqstrToIdent (s : String) : String := Id.run <| do 32 | let mut ret := "" 33 | let mut curr : String.Pos := ⟨0⟩ 34 | let mut sqcnt := 0 35 | while true do 36 | match s.get? curr with 37 | | some ch => 38 | if ch == '\'' then 39 | if sqcnt == 0 then 40 | ret := ret.push '«' 41 | else 42 | ret := ret.push '»' 43 | sqcnt := (sqcnt + 1) % 2 44 | else 45 | ret := ret.push ch 46 | curr := curr + ch 47 | | none => break 48 | return ret 49 | 50 | def splitOnOutermostPeriod (s : String) : Array String := Id.run <| do 51 | let mut ret := #[] 52 | let mut last : String.Pos := ⟨0⟩ 53 | let mut curr : String.Pos := ⟨0⟩ 54 | let mut depth := 0 55 | while true do 56 | match s.get? curr with 57 | | some ch => 58 | curr := curr + ch 59 | if ch == '(' then 60 | depth := depth + 1 61 | if ch == ')' then 62 | depth := depth - 1 63 | if ch == '.' && depth == 0 then 64 | ret := ret.push (s.extract last curr) 65 | last := curr 66 | | none => break 67 | return ret 68 | 69 | def loadTptp (path : System.FilePath) : CommandElabM (Syntax × Nat) := do 70 | let lines ← IO.FS.lines path 71 | let lines := lines.filter fun l => ¬ l.startsWith "%" 72 | let s := String.join lines.toList 73 | -- Replace `$` with `🍉` so that it won't conflict with Lean's antiquot 74 | let s := s.replace "$" "🍉" 75 | let sarr := (splitOnOutermostPeriod s).map sqstrToIdent 76 | let mut stxarr : Array (TSyntax `TPTP_file) := #[] 77 | -- Parse input-by-input so that the parser is easier to debug 78 | for s in sarr do 79 | stxarr := stxarr.push ⟨← parseTPTPInput s⟩ 80 | return (← `(TPTP_file| $[$stxarr]*), lines.size) 81 | 82 | partial def resolveInclude (leadingPath : System.FilePath) : Syntax → CommandElabM (Syntax × Nat) 83 | |`(TPTP_file| $[$f]*) => do 84 | let mut result := #[] 85 | let mut lines := 0 86 | for stx in f do 87 | let (stx, lineno) ← resolveInclude leadingPath stx 88 | lines := lines + lineno 89 | match stx with 90 | |`(TPTP_file| $[$g]*) => result := result.append g 91 | |`(TPTP_input| include( $_ ).) => throwError "resolveInclude :: include is not resolved in {stx}" 92 | | other => result := result.push other 93 | let stx ← `(TPTP_file| $[$result]*) 94 | return (stx, lines) 95 | |`(TPTP_input| include( $ri ).) => do 96 | let path := leadingPath / (Lean.Syntax.getId ri.raw).getString! 97 | loadTptp path 98 | | other => return (other, 0) 99 | 100 | syntax (name := tptpKind) "tptp " ident strLit term : command 101 | 102 | @[command_elab tptpKind] def elabResolve : CommandElab := fun stx => do 103 | match stx with 104 | | `(tptp $name $file $proof) => 105 | match Syntax.isStrLit? file with 106 | | some file => 107 | let (fstx, lines) ← loadTptp file 108 | let components := (⟨file⟩ : System.FilePath).components 109 | let leadingPath := System.mkFilePath (components.take (components.length - 3)) 110 | let (fstxResolved, extraLines) ← resolveInclude leadingPath fstx 111 | checkMaxTPTPProblemLines (lines + extraLines) 112 | elabCommand (← `(BEGIN_TPTP $name $fstxResolved END_TPTP $proof)) 113 | | _ => throwError "Expected strLit: {file}" 114 | | _ => throwError "Failed to parse tptp command" -------------------------------------------------------------------------------- /Duper/TPTPParser/SyntaxDecl.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | 3 | open Lean 4 | open Lean.Parser 5 | open TSyntax.Compat 6 | 7 | namespace TPTP 8 | 9 | declare_syntax_cat TPTP_file 10 | 11 | declare_syntax_cat thf_type 12 | declare_syntax_cat thf_term 13 | declare_syntax_cat thf_atomic_type 14 | 15 | syntax thf_arguments := "(" thf_term,* ")" 16 | syntax:120 rawIdent thf_arguments ? : thf_term 17 | -- TODO : Add support for float? 18 | syntax:max num : thf_term 19 | syntax:max "(" thf_term ")" : thf_term 20 | 21 | -- Higher-order Application 22 | syntax defined_term := "🍉" noWs ident 23 | syntax:max defined_term : thf_term 24 | syntax:60 thf_term:60 "@" thf_term:61 : thf_term 25 | syntax bexpOp := "|" <|> "&" <|> "<=>" <|> "=>" <|> "<=" <|> "<~>" <|> "~|" <|> "~&" 26 | syntax:60 thf_term:60 bexpOp thf_term:61 : thf_term 27 | -- Approximation of `thf_binary_type` 28 | syntax:60 thf_term:61 ">" thf_term:60 : thf_term 29 | 30 | -- ::= = 31 | -- ::= != 32 | -- ::= 33 | -- ::= 34 | -- 35 | -- ::= 36 | -- 37 | 38 | syntax eqOp := "=" <|> "!=" 39 | syntax:75 thf_term:90 eqOp thf_term:90 : thf_term 40 | syntax:70 "~" thf_term:70 : thf_term 41 | 42 | syntax annotation := "," rawIdent 43 | 44 | syntax defined_type := "🍉" noWs ident 45 | syntax:max defined_type : thf_atomic_type 46 | syntax rawIdent : thf_atomic_type 47 | syntax thf_atomic_type : thf_type 48 | syntax:max "(" thf_type ")" : thf_type 49 | 50 | def thfXProdArgsParser := sepBy1 (categoryParser `thf_atomic_type 0) "*" 51 | syntax thf_xprod_args := thfXProdArgsParser 52 | 53 | --thf_mapping_type 54 | /- 55 | Note: In real TPTP syntax, the below line should be: tff_atomic_type > tff_atomic_type : thf_type. 56 | However, this is infeasible because Lean's parser can't reliably identify thf_atomic_types due 57 | to an issue with how Lean's built-in antiquotations and our defined_type category conflict with each 58 | other. So writing the below syntax is a workaround. 59 | -/ 60 | syntax:120 thf_type:121 ">" thf_type:120 : thf_type 61 | syntax "(" thf_xprod_args ")" ">" thf_atomic_type : thf_type 62 | 63 | --thf_type_arguments are needed because () is a thf_atomic_type 64 | syntax thf_type_arguments := "(" thf_atomic_type,* ")" 65 | syntax rawIdent thf_type_arguments : thf_atomic_type 66 | 67 | --thf_apply_type 68 | syntax:130 thf_type:130 "@" thf_type:131 : thf_type 69 | 70 | syntax quantifier := "!" <|> "?" <|> "^" 71 | syntax thf_variable := rawIdent (":" thf_type) ? 72 | syntax:70 quantifier "[" thf_variable,* "]" ":" thf_term:70 : thf_term 73 | 74 | syntax th1_quantifier := "!>" 75 | syntax th1_quantifier "[" thf_variable,* "]" ":" thf_type : thf_type --tf1_quantified_type 76 | 77 | declare_syntax_cat TPTP_input 78 | syntax "tff" "(" rawIdent "," rawIdent "," thf_term annotation ? ")" "." : TPTP_input 79 | syntax "tff" "(" rawIdent "," &"type" "," rawIdent ":" thf_type annotation ? ")" "." : TPTP_input 80 | syntax "thf" "(" rawIdent "," rawIdent "," thf_term annotation ? ")" "." : TPTP_input 81 | syntax "thf" "(" rawIdent "," &"type" "," rawIdent ":" thf_type annotation ? ")" "." : TPTP_input 82 | syntax "cnf" "(" rawIdent "," rawIdent "," thf_term annotation ? ")" "." : TPTP_input 83 | syntax "fof" "(" rawIdent "," rawIdent "," thf_term annotation ? ")" "." : TPTP_input 84 | syntax "include" "(" rawIdent ")" "." : TPTP_input 85 | 86 | syntax TPTP_input * : TPTP_file -------------------------------------------------------------------------------- /Duper/Tests/DApplyTests.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import Duper.DUnif.DApply 3 | 4 | -- whnf test 5 | def wh₁ : Nat := 3 6 | 7 | set_option dUnifDbgOn true 8 | 9 | set_option trace.Meta.Tactic true in 10 | def wh₀ (f : Nat → Prop) (h : ∀ x, f x) : f wh₁ := 11 | by dapply h attempt 5 unifier 0 contains 0 12 | 13 | -- Imitation 14 | set_option trace.duper.dUnif.debug true in 15 | def imt₀ (f : Nat → Prop) (h : ∀ x, f x) : f 3 := 16 | by dapply h attempt 30 unifier 0 contains 0 17 | 18 | def imt₁ (f : Nat → Prop) 19 | (h : ∀ x y (z : Nat), f z → f (x + y)) 20 | (k : f 0) : f (3 + b) := by 21 | dapply h attempt 30 unifier 0 contains 0 22 | case a => dapply h attempt 70 unifier 0 contains 0; 23 | case a => apply k; 24 | exact 0; exact 0 25 | 26 | def imt₂ (f : Nat → Prop) 27 | (comm : ∀ x y, f (x + y) → f (y + x)) 28 | (h : ∀ x y z, f (x + z) → f (x + y)) 29 | (g : f (u + v)) 30 | : f (a + b) := by 31 | dapply h attempt 30 unifier 0 contains 0 32 | case a => 33 | dapply comm attempt 42 unifier 0 contains 0; 34 | case a => dapply h attempt 42 unifier 0 contains 0; 35 | case a => dapply g attempt 42 unifier 0 contains 0 36 | 37 | -- Huet-Style Projection 38 | def hsp₁ (ftr : ∀ (f g : Nat → Prop) (x : Nat), f x → g x) 39 | (S T : Nat → Prop) (u v : Nat) 40 | (h : S u) : T v := by 41 | dapply ftr attempt 100 unifier 1 contains 0 42 | case a => dapply h attempt 10 unifier 0 contains 0 43 | 44 | def hsp₂ (p : Nat → Prop) (x y : Nat) 45 | (hp : ∀ (f : Nat → Nat → Nat), p (f x y) ∧ p (f y x)) 46 | : p x ∧ p y := by 47 | dapply hp attempt 11 unifier 0 contains 0 48 | 49 | def hsp₃ (p : Nat → Prop) (x y : Nat) 50 | (hp : ∀ (f : Nat → Nat → Nat), p (f x y) ∧ p (f y x)) 51 | : p (x + y) ∧ p (y + x) := by 52 | dapply hp attempt 300 unifier 0 contains 0 53 | 54 | def hsp₄ (done : Prop) 55 | (gene : ∀ (H : (Nat → Nat) → Nat → Nat), (fun F X => H F X) = (fun F X => F X) → done) : 56 | done := by 57 | dapply gene attempt 10 unifier 0 contains 0 58 | case a => dapply Eq.refl attempt 70 unifier 0 contains 0; 59 | 60 | opaque www : Nat → Nat → Nat := fun _ _ => 1 61 | opaque ww : Nat → Nat := id 62 | opaque w : Nat 63 | 64 | -- Elimination 65 | def elm₁ (p : Nat → Prop) 66 | (a b : Nat) (h : ∀ (f : Nat → Nat → Nat), p (f a b)) 67 | (g : ∀ (ay : Nat), p ay → False) 68 | : False := by 69 | dapply g attempt 10 unifier 0 contains 0 70 | case a => dapply h attempt 590 unifier 0 contains 0; exact www 71 | 72 | #print elm₁.proof_1 73 | 74 | -- Identification 75 | def idt₁ (p : Nat → Prop) (x : Nat) 76 | (hp : ∀ (f g : Nat → Nat), p (f (g x)) ∧ p (g (f x))) 77 | (done : Prop) 78 | (hq : ∀ w, p w ∧ p w → done) : done := by 79 | dapply hq attempt 10 unifier 0 contains 0 80 | case a => dapply hp attempt 10000 unifier 66 contains 0 81 | exact id; exact (fun _ => (Nat → Nat)); exact (fun _ b _ => b id) 82 | 83 | #print idt₁.proof_1 84 | 85 | -- Negative tests 86 | def neg₁ (h : True) : False := 87 | by dapply h attempt 10 unifier 0 contains 0 88 | -------------------------------------------------------------------------------- /Duper/Tests/bugs.lean: -------------------------------------------------------------------------------- 1 | import Duper.TPTP 2 | import Duper.Tactic 3 | 4 | -- Bug 1 5 | namespace Color2 6 | 7 | inductive Color := 8 | | red : Color 9 | | green : Color 10 | 11 | example : @Color.rec (fun _ => Nat) a b .red = a := by duper [Color.rec] 12 | 13 | def test : Color → Color 14 | | .red => .green 15 | | .green => .red 16 | 17 | set_option pp.match false 18 | #print test 19 | #print test.match_1 20 | #print Color.casesOn 21 | 22 | -- This bug demonstrates the difficulty Duper has dealing with inductive arguments and facts such as Color.rec 23 | set_option inhabitationReasoning false in 24 | set_option trace.duper.saturate.debug true in 25 | set_option trace.duper.timeout.debug true in 26 | example : ∃ c : Color, test c = .green := by 27 | duper [test, test.match_1, Color.rec, Color.casesOn] {portfolioInstance := 0} 28 | 29 | example : Color.red ≠ Color.green := by 30 | duper {portfolioInstance := 0} 31 | /- 32 | It seems that reasoning about inductive types requires solving unification problems 33 | that are completely out of the scope of the current algorithms duper uses 34 | Here is an example of proving `Color.red ≠ Color.green` without using extra 35 | inductive types like `Color.noConfusion` 36 | -/ 37 | theorem red_ne_green : Color.red ≠ Color.green := fun h => 38 | let color_case := fun (t : Color) => @Color.casesOn (fun _ => Prop) t True False 39 | have color_true : color_case Color.red := True.intro 40 | (h ▸ color_true : color_case Color.green) 41 | /- 42 | What we're trying to here do is to rewrite `True` to `False` using `Color.red = Color.green` 43 | To do this, notice that `True ≝ color_case Color.red` and `False ≝ color_case Color.green` 44 | -/ 45 | 46 | end Color2 47 | 48 | -- What was previously bug 5 (though I would no longer call it a bug per se) 49 | example : ((∃ (A B : Type) (f : B → A) (x : B), f x = f x) = True) := 50 | by duper {portfolioInstance := 7} 51 | 52 | example : ((∃ (A B : Type) (f : B → A) (x : B), f x = f x) = True) := 53 | by duper {portfolioInstance := 8} 54 | 55 | /- 56 | The previous bug 5 error of unknown fvars has been corrected. However, this example still fails when 57 | inhabitation reasoning is enabled. The issue that occurs is that when inhabitation reasoning is enabled, 58 | the clause `∀ (a a_1 : Type), (a_1 → a) → a_1 → False` is recognized as potentially vacuous. This makes 59 | sense because if `a_1` is assigned an empty type, then no contradiction can be yielded by this clause. 60 | 61 | Duper happens to be able to derive a contradiction from this clause when inhabitation reasoning is disabled 62 | because the default instances that `Lean.Meta.findInstance` attempts happen to work well for this example, 63 | though there are other examples where that strategy would not work. For instance, Duper will fail to 64 | derive a contradiction from the clause `∀ x : Nat, Fin x → False` because the default `Nat` is 0. If the 65 | default Nat were 1, then Duper would be able to derive a contradiction from this when inhabitation reasoning 66 | is disabled (though Duper would still consider the clause vacuous with inhabitation reasoning enabled). 67 | 68 | I wouldn't call the current behavior a bug so much as an area where Duper has a lot of room for improvement. 69 | Supporting dependent type reasoning is a relatively low priority, but I'll leave this example documented 70 | here for the future. More examples of related behavior are included at the end of test_inhabitationReasoning.lean. 71 | -/ 72 | 73 | -- Diagnosis of the above test 74 | /- 75 | The error appears to occur during removeVanishedVarsHelper (which is only called when inhabitationReasoning is enabled). 76 | -/ 77 | 78 | -- Bug 8: Application type mismatch pertaining to the universe levels of skolems 79 | axiom f.{u} : Type u → Prop 80 | 81 | axiom ftrue.{u} : f.{u} (Sort u) 82 | 83 | axiom exftrue.{u} : ∃ (x : Type u), f x 84 | set_option trace.duper.proofReconstruction true in 85 | def test : ∃ x : Type u, ∃ y : Type v, f x = f y := by 86 | duper [exftrue] 87 | /- 88 | application type mismatch 89 | skol.0 [anonymous] 90 | argument has type 91 | Type u_1 92 | but function has type 93 | Type u_2 → Type u_2 94 | -/ 95 | 96 | -- Bug 9 97 | -- This example can only be solved when inhabitationReasoning is disabled (saturates if inhabitationReasoning is enabled) 98 | set_option inhabitationReasoning false in 99 | example : ∃ (A : Type) (B : A → Type) (f : ∀ (a : A), B a) (x : A), (f x = f x) = True := 100 | by duper {portfolioInstance := 0} 101 | 102 | set_option inhabitationReasoning true in 103 | example : ∃ (A : Type) (B : A → Type) (f : ∀ (a : A), B a) (x : A), (f x = f x) = True := 104 | by duper {portfolioInstance := 0} 105 | -------------------------------------------------------------------------------- /Duper/Tests/fail_group.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | 3 | -- vampire succeeds on this example 4 | 5 | theorem cubeisom 6 | (G : Type) 7 | (e : G) 8 | (mult : G → G → G) 9 | (inverse : G → G) 10 | (cube : G → G) 11 | (left_identity : ∀ x, mult e x = x) 12 | (left_inverse : ∀ x, mult (inverse x) x = e) 13 | (associativity : ∀ x y z, mult (mult x y) z = mult x (mult y z)) 14 | (cube_definition : ∀ x, cube x = mult x (mult x x)) 15 | (cube_injective : ∀ x y, cube x = cube y → x = y) 16 | (cube_surjective : ∀ x, ∃ y, x = cube y) 17 | (cube_homomorphism : ∀ x y, cube (mult x y) = mult (cube x) (cube y)) 18 | : ∀ x y, mult x y = mult y x := by duper [*] 19 | -------------------------------------------------------------------------------- /Duper/Tests/fail_group.tptp: -------------------------------------------------------------------------------- 1 | fof(left_identity,axiom, 2 | ! [X] : mult(e,X) = X). 3 | 4 | fof(left_inverse,axiom, 5 | ! [X] : mult(inverse(X),X) = e). 6 | 7 | fof(associativity,axiom, 8 | ! [X,Y,Z] : mult(mult(X,Y),Z) = mult(X,mult(Y,Z))). 9 | 10 | fof(cube_definition,hypothesis, 11 | ! [X] : cube(X) = mult(X, mult(X, X))). 12 | 13 | fof(cube_injective,hypothesis, 14 | ! [X,Y] : cube(X) = cube(Y) => X = Y). 15 | 16 | fof(cube_surjective,hypothesis, 17 | ! [X] : ? [Y] : X = cube(Y)). 18 | 19 | fof(cube_homomorphism,hypothesis, 20 | ! [X,Y] : cube(mult(X,Y)) = mult(cube(X),cube(Y))). 21 | 22 | fof(commutivity,conjecture, 23 | ! [X,Y] : mult(X,Y) = mult(Y,X)). 24 | 25 | %---cmd: ./vampire_z3_Release_static_master_4764 -m 12000 -t 1h Group_Theory/CubeIsom.tptp --mode casc -------------------------------------------------------------------------------- /Duper/Tests/fail_lattice.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | 3 | axiom L : Type 4 | axiom U : L → L → L 5 | axiom A : L → L → L 6 | axiom LLE : L → L → Prop 7 | infix:90 "⋂" => A 8 | infix:90 "⋃" => U 9 | 10 | instance : LE L where 11 | le a b := a = a ⋂ b 12 | 13 | -- vampire succeeds on this example 14 | 15 | theorem mod_lattice1 16 | (U_Comm : ∀ a b, U a b = U b a) 17 | (A_Comm : ∀ a b, A a b = A b a) 18 | (U_Assoc : ∀ a b c, U a (U b c) = U (U a b) c) 19 | (A_Assoc : ∀ a b c, A a (A b c) = A (A a b) c) 20 | (L_Lattice_U : ∀ a b, U a (A a b) = a) 21 | (L_Lattice_A : ∀ a b, A a (U a b) = a) 22 | (ModLatA : ∀ a x b, U (A a b) (A x b) = A (U (A a b) x) b) 23 | (a b c : L) 24 | (Hyp : A a (U b c) = U (A a b) (A a c)) 25 | : U a (A b c) = A (U a b) (U a c) := by duper [*] -------------------------------------------------------------------------------- /Duper/Tests/fail_lattice.tptp: -------------------------------------------------------------------------------- 1 | fof(union_comm,axiom, 2 | ! [A,B] : union(A,B) = union(B,A)). 3 | 4 | fof(inter_comm,axiom, 5 | ! [A,B] : inter(A,B) = inter(B,A)). 6 | 7 | fof(union_assoc,axiom, 8 | ! [A,B,C] : union(A,union(B,C)) = union(union(A,B),C)). 9 | 10 | fof(inter_assoc,axiom, 11 | ! [A,B,C] : inter(A,inter(B,C)) = inter(inter(A,B),C)). 12 | 13 | fof(lattice_U,axiom, 14 | ! [A,B] : union(A,inter(A,B)) = A). 15 | 16 | fof(lattice_A,axiom, 17 | ! [A,B] : inter(A,union(A,B)) = A). 18 | 19 | fof(modLat_A,axiom, 20 | ! [A,X,B] : union(inter(A,B),inter(X,B)) = inter(union(inter(A,B),X),B)). 21 | 22 | fof(hyp,hypothesis, 23 | inter(a,union(b,c)) = union(inter(a,b),inter(a,c))). 24 | 25 | fof(conj,conjecture, 26 | union(a,inter(b,c)) = inter(union(a,b),union(a,c))). 27 | 28 | %---cmd: ./vampire_z3_Release_static_master_4764 -m 12000 -t 1h Group_Theory/ModLattice.tptp --mode casc -------------------------------------------------------------------------------- /Duper/Tests/fail_tests.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | import Duper.TPTP 3 | 4 | -- Tests where the prover should saturate 5 | 6 | axiom a : Nat 7 | axiom b : Nat 8 | axiom c : Nat 9 | axiom d : Nat 10 | axiom f : Nat → Nat 11 | 12 | theorem test0018saturate (a1 a2 a3 a4 a5 a6 : Nat) 13 | (h1 : 14 | f (f (f (f (f (f (f (f a5))))))) = d ∨ 15 | f (f (f (f (f (f (f a4)))))) = d ∨ 16 | f (f (f (f (f (f a3))))) = d ∨ 17 | f (f (f (f (f a2)))) = d ∨ 18 | f (f (f (f a1))) = d ∨ 19 | f (f (f a)) = d ∨ f (f b) = d ∨ f c = d) 20 | (h2 : f (f (f (f (f (f (f (f a5))))))) ≠ d) 21 | (h2 : f (f (f (f (f (f (f a4)))))) ≠ d) 22 | (h2 : f (f (f (f (f (f a3))))) ≠ d) 23 | (h2 : f (f (f (f (f a2)))) ≠ d) 24 | (h2 : f (f (f (f a1))) ≠ d) 25 | (h2 : f (f (f a)) ≠ d) 26 | (h3 : f (f b) ≠ d) 27 | : False := by duper [*] 28 | 29 | -- Tests where duper should time out 30 | -- This example is intended to test duper's ability of 31 | -- handling dependent types. It should fail with 32 | -- "deterministic timeout". 33 | -- If it fails in any other way, it's an indication of a bug. 34 | 35 | set_option trace.Meta.debug true in 36 | def rec₁ : False := by 37 | duper [Nat.rec] -------------------------------------------------------------------------------- /Duper/Tests/ffffa.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | 3 | axiom f : Nat → Nat 4 | axiom a : Nat 5 | 6 | example (h : f a = a) : 7 | f (f (f (f (f (f (f (f (f (f ( 8 | f (f (f (f (f (f (f (f (f (f ( 9 | f (f (f (f (f (f (f (f (f (f ( 10 | f (f (f (f (f (f (f (f (f (f ( 11 | f (f (f (f (f (f (f (f (f (f ( 12 | a 13 | )))))))))) 14 | )))))))))) 15 | )))))))))) 16 | )))))))))) 17 | )))))))))) = a 18 | := by duper [h] -------------------------------------------------------------------------------- /Duper/Tests/fluidSupTests.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | import Duper.TPTP 3 | 4 | set_option trace.duper.saturate.debug true 5 | set_option trace.duper.rule.fluidSup true 6 | 7 | -- Examples taken from https://matryoshka-project.github.io/pubs/lamsup_article.pdf pages 9 and 10 8 | 9 | theorem supWithLambdasEx13 (a b c : t) (f g : t → t) (h : t → t → t) 10 | (eq1 : f a = c) (eq2 : ∀ y : t → t, h (y b) (y a) ≠ h (g (f b)) (g c)) : False := 11 | by duper [*] {portfolioInstance := 0} 12 | 13 | theorem supWithLambdasEx14 (a b c d : t) (f g : t → t) 14 | (eq1 : f a = c) (eq2 : f b = d) (eq3 : ∀ y : t → t, g c ≠ y a ∨ g d ≠ y b) : False := 15 | by duper [*] 16 | 17 | -- Note: Since the original example was untyped, it's possible that the types I chose for this don't work. So I don't 18 | -- put too much stock in it if this continues to fail. But I do think that supWithLambda13 should be solvable 19 | theorem supWithLambdaEx15 (a c : t) (f g : t → t) (h : t → ((t → t) → t → t) → t) 20 | (eq1 : f a = c) (eq2 : ∀ y : (t → t) → t → t, h (y (fun x => g (f x)) a) y ≠ h (g c) (fun w x => w x)) : 21 | False := by duper [*] 22 | -------------------------------------------------------------------------------- /Duper/Tests/group.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | 3 | axiom G : Type 4 | axiom e : G 5 | axiom mul : G → G → G 6 | 7 | infix:80 (priority := high) " ⬝ " => mul 8 | 9 | axiom mul_assoc : ∀ (x y z : G), (x ⬝ y) ⬝ z = x ⬝ (y ⬝ z) 10 | axiom mul_e : ∀ (x : G), x ⬝ e = x 11 | axiom exists_left_inv : ∀ (x : G), ∃ (y : G), x ⬝ y = e 12 | 13 | noncomputable instance : Inhabited G := ⟨e⟩ 14 | 15 | theorem e_mul : e ⬝ x = x := 16 | by duper [mul_assoc, mul_e, exists_left_inv] {portfolioInstance := 1} 17 | 18 | theorem exists_right_inv (x : G) : ∃ (y : G), y ⬝ x = e := 19 | by duper [mul_assoc, mul_e, exists_left_inv] {portfolioInstance := 1} 20 | 21 | theorem left_neutral_unique (x : G) : (∀ y, x ⬝ y = y) → x = e := 22 | by duper [mul_assoc, mul_e, exists_left_inv] {portfolioInstance := 1} 23 | 24 | theorem right_neutral_unique (x : G) : (∀ y, y ⬝ x = y) → x = e := 25 | by duper [mul_assoc, mul_e, exists_left_inv] {portfolioInstance := 1} 26 | 27 | theorem right_inv_unique (x y z : G) (h : x ⬝ y = e) (h : x ⬝ z = e) : y = z := 28 | by duper [*, mul_assoc, mul_e, exists_left_inv] {portfolioInstance := 1} 29 | 30 | theorem left_inv_unique (x y z : G) (h : y ⬝ x = e) (h : z ⬝ x = e) : y = z := 31 | by duper [*, mul_assoc, mul_e, exists_left_inv] {portfolioInstance := 1} 32 | 33 | noncomputable def sq (x : G) := x ⬝ x 34 | 35 | theorem sq_mul_sq_eq_e (x y : G) (h_comm : ∀ a b, a ⬝ b = b ⬝ a) (h : x ⬝ y = e) : 36 | sq x ⬝ sq y = e := 37 | by duper [h_comm, h, sq, mul_assoc, mul_e] {portfolioInstance := 1} 38 | -------------------------------------------------------------------------------- /Duper/Tests/group2.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | 3 | class Group (G : Type) where 4 | (one : G) 5 | (inv : G → G) 6 | (mul : G → G → G) 7 | (mul_assoc : ∀ (x y z : G), mul (mul x y) z = mul x (mul y z)) 8 | (mul_one : ∀ (x : G), mul x one = x) 9 | (mul_inv : ∀ (x : G), mul x (inv x) = one) 10 | 11 | namespace Group 12 | 13 | variable {G : Type} [hG : Group G] (x y : G) 14 | 15 | infix:80 (priority := high) " ⬝ " => Group.mul 16 | 17 | noncomputable instance : Inhabited G := ⟨one⟩ 18 | 19 | theorem test : x ⬝ one = x := 20 | by duper [Group.mul_one] {portfolioInstance := 1} 21 | 22 | theorem exists_right_inv (x : G) : inv x ⬝ x = one := 23 | by duper [Group.mul_assoc, Group.mul_one, Group.mul_inv] {portfolioInstance := 1} 24 | 25 | theorem left_neutral_unique (x : G) : (∀ y, x ⬝ y = y) → x = one := 26 | by duper [Group.mul_assoc, Group.mul_one, Group.mul_inv] {portfolioInstance := 1} 27 | 28 | theorem right_neutral_unique (x : G) : (∀ y, y ⬝ x = y) → x = one := 29 | by duper [Group.mul_assoc, Group.mul_one, Group.mul_inv] {portfolioInstance := 1} 30 | 31 | theorem right_inv_unique (x y z : G) (h : x ⬝ y = one) (h : x ⬝ z = one) : y = z := 32 | by duper [*, Group.mul_assoc, Group.mul_one, Group.mul_inv] {portfolioInstance := 1} 33 | 34 | theorem left_inv_unique (x y z : G) (h : y ⬝ x = one) (h : z ⬝ x = one) : y = z := 35 | by duper [*, Group.mul_assoc, Group.mul_one, Group.mul_inv] {portfolioInstance := 1} 36 | 37 | noncomputable def sq := x ⬝ x 38 | 39 | theorem sq_mul_sq_eq_e (h_comm : ∀ (a b : G), a ⬝ b = b ⬝ a) (h : x ⬝ y = one) : 40 | sq x ⬝ sq y = one := 41 | by duper [h, h_comm, sq, Group.mul_assoc, Group.mul_one, Group.mul_inv] {portfolioInstance := 1} 42 | 43 | end Group 44 | -------------------------------------------------------------------------------- /Duper/Tests/lastAsylum.lean: -------------------------------------------------------------------------------- 1 | import Duper.TPTP 2 | import Duper.Tactic 3 | 4 | -- inhabitants 5 | axiom Inhab : Type 6 | 7 | axiom Fether : Inhab 8 | axiom Tarr : Inhab 9 | axiom Jones : Inhab 10 | axiom Smith : Inhab 11 | axiom A : Inhab 12 | axiom B : Inhab 13 | axiom C : Inhab 14 | axiom D : Inhab 15 | axiom Sane : Inhab → Prop 16 | axiom Doctor : Inhab → Prop 17 | axiom Peculiar : Inhab → Prop 18 | axiom Special : Inhab → Prop 19 | axiom bf : Inhab → Inhab 20 | 21 | axiom ax1 : ∀ x, Peculiar x ↔ (Sane x ↔ ¬ Doctor x) 22 | axiom ax2 : ∀ x, Special x ↔ ∀ y, (¬ Doctor y ↔ (Sane y ↔ Peculiar x)) 23 | axiom ax3 : ∀ x y, (Sane x ↔ Special y) → (Sane (bf x) ↔ ¬ Doctor y) 24 | axiom ax4 : Sane Tarr ↔ ∀ x, Doctor x → Sane x 25 | axiom ax5 : Sane Fether ↔ ∀ x, ¬ Doctor x → ¬ (Sane x) 26 | axiom ax6 : Sane Fether ↔ Sane Tarr 27 | 28 | example 29 | (ax1 : ∀ x, Peculiar x ↔ (Sane x ↔ ¬ Doctor x)) 30 | (ax2 : ∀ x, Special x ↔ ∀ y, (¬ Doctor y ↔ (Sane y ↔ Peculiar x))) 31 | (ax3 : ∀ x y, (Sane x ↔ Special y) → (Sane (bf x) ↔ ¬ Doctor y)) 32 | (ax4 : Sane Tarr ↔ ∀ x, Doctor x → Sane x) 33 | (ax5 : Sane Fether ↔ ∀ x, ¬ Doctor x → ¬ (Sane x)) 34 | (ax6 : Sane Fether ↔ Sane Tarr) 35 | : False := sorry --by duper -- timeout 36 | 37 | theorem asylum_one 38 | (h1 : Sane Jones ↔ Doctor Smith) 39 | (h2 : Sane Smith ↔ ¬ Doctor Jones) 40 | : ∃ x : Inhab, (¬ Sane (x) ∧ Doctor (x)) ∨ (Sane (x) ∧ ¬ Doctor (x)) := by duper [h1, h2] 41 | 42 | #print axioms asylum_one 43 | 44 | theorem asylum_seven 45 | (h1 : Sane A ↔ ¬ Sane B) 46 | (h2 : Sane B ↔ Doctor A) 47 | : (¬ Sane A ∧ Doctor A) ∨ (Sane A ∧ ¬ Doctor A) := by duper [*] 48 | 49 | #print axioms asylum_seven 50 | 51 | set_option trace.duper.timeout.debug true 52 | 53 | theorem asylum_nine 54 | (h1 : Sane A ↔ (Sane B ↔ Sane C)) 55 | (h2 : Sane B ↔ (Sane A ↔ Sane D)) 56 | (h3 : Sane C ↔ ¬ (Doctor C ∧ Doctor D)) 57 | (h4 : A ≠ B ∧ A ≠ C ∧ A ≠ D ∧ B ≠ C ∧ B ≠ D ∧ C ≠ D) : 58 | (∃ x : Inhab, Sane x ∧ ¬ Doctor x) ∨ 59 | (∃ x : Inhab, ∃ y : Inhab, x ≠ y ∧ (¬ Sane x) ∧ Doctor x ∧ (¬ Sane y) ∧ Doctor y) := 60 | by duper [*] 61 | 62 | #print axioms asylum_nine 63 | #print asylum_nine 64 | -------------------------------------------------------------------------------- /Duper/Tests/monobugs.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | import Duper.TPTP 3 | import Auto.Tactic 4 | 5 | -- This file is for documenting bugs in Duper's current invocation of the monomorphization procedure. 6 | 7 | -- Note: This example fails if inhabitation reasoning is enabled (see bug 5 in bugs.lean) or if monomorphization is enabled 8 | example : ((∃ (A B : Type) (f : B → A) (x : B), f x = f x) = True) := 9 | by duper {portfolioInstance := 6, inhabitationReasoning := false} 10 | 11 | example : ((∃ (A B : Type) (f : B → A) (x : B), f x = f x) = True) := 12 | by duper {portfolioInstance := 1, inhabitationReasoning := false} 13 | 14 | -- allNative_direct :: Cannot find external proof of nonempty #0 15 | tptp SEU123 "../TPTP-v8.0.0/Problems/SEU/SEU123+1.p" 16 | by duper [*] {portfolioInstance := 6} 17 | 18 | tptp SEU123_new "../TPTP-v8.0.0/Problems/SEU/SEU123+1.p" 19 | by duper [*] {portfolioInstance := 1} 20 | 21 | -- unexpected bound variable #0 22 | theorem boolSimpRule27TestDep₁ (a b c y z r : Prop) (f : a ∧ b ∧ c → Prop) (h : ((x : a ∧ b ∧ c) → (y ∨ f x ∨ c ∨ z)) = r) : r := 23 | by duper [*] {portfolioInstance := 6} 24 | 25 | theorem boolSimpRule27TestDep₁_new (a b c y z r : Prop) (f : a ∧ b ∧ c → Prop) (h : ((x : a ∧ b ∧ c) → (y ∨ f x ∨ c ∨ z)) = r) : r := 26 | by duper [*] {portfolioInstance := 1} 27 | 28 | namespace NegativeBoolSimpTests 29 | 30 | axiom f.{u} : Sort u → Nat 31 | 32 | -- reifTermCheckType :: LamTerm (¬ ((!0 (∀ x0 : Nat, !1)) = (!0 (∀ x0 : Nat, !1)))) is not type correct 33 | def neg₁ : (f (Nat → Nat) = f (Nat → Nat)) := by duper {portfolioInstance := 6} 34 | 35 | def neg₁_new : (f (Nat → Nat) = f (Nat → Nat)) := by duper {portfolioInstance := 1} 36 | 37 | end NegativeBoolSimpTests 38 | 39 | axiom f.{u} : Type u → Prop 40 | 41 | axiom ftrue.{u} : f.{u} (Sort u) 42 | 43 | -- prover saturated 44 | def unitst₁ : f.{max u v} (Sort (max u v)) ∧ f.{v} (Sort v) := by 45 | duper [ftrue] {portfolioInstance := 6} 46 | 47 | def unitst₁_new : f.{max u v} (Sort (max u v)) ∧ f.{v} (Sort v) := by 48 | duper [ftrue] {portfolioInstance := 1} 49 | 50 | axiom fmoretrue.{u} : ∀ (x : Type u), f x 51 | 52 | -- prover saturated 53 | def unitst₂ : ∀ (x : Type v), f x := by 54 | duper [fmoretrue] {portfolioInstance := 6} 55 | 56 | def unitst₂_new : ∀ (x : Type v), f x := by 57 | duper [fmoretrue] {portfolioInstance := 1} 58 | 59 | axiom exftrue.{u} : ∃ (x : Type u), f x 60 | 61 | -- prover saturated 62 | def skuniverse.{u} : ∃ (x : Type u), f x := by 63 | duper [exftrue] {portfolioInstance := 6} 64 | 65 | def skuniverse_new.{u} : ∃ (x : Type u), f x := by 66 | duper [exftrue] {portfolioInstance := 1} 67 | 68 | -- prover saturated (presumably because monomorphization procedure doesn't use nonempty lemmas in the same way duper does) 69 | example (h : Nonempty (α → β)) : (∀ n : Nat, ∀ a : α, ∃ b : β, True) = True := 70 | by duper [h] {portfolioInstance := 6} 71 | 72 | example (h : Nonempty (α → β)) : (∀ n : Nat, ∀ a : α, ∃ b : β, True) = True := 73 | by duper [h] {portfolioInstance := 1} 74 | -------------------------------------------------------------------------------- /Duper/Tests/profiling.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | 3 | open Lean Core Meta 4 | open Elab.Tactic 5 | 6 | -- Trying to see whether monad stack is more efficient than 7 | -- building two monads on top of the same base monad. 8 | namespace MonadStackPerformance 9 | 10 | structure State1 where 11 | mctx : MetavarContext := {} 12 | lctx : LocalContext := {} 13 | extra : Nat := 0 14 | 15 | abbrev RuleM1 := StateRefT State1 CoreM 16 | 17 | def runMetaAsRuleM1 (m : MetaM α) : RuleM1 α := do 18 | let lctx := (← get).lctx 19 | let mctx := (← get).mctx 20 | let (ret, state) ← Meta.MetaM.run (ctx := {lctx := lctx}) (s := {mctx := mctx}) m 21 | modify (fun s => {s with mctx := state.mctx}) 22 | return ret 23 | 24 | def ruleM1InferType : RuleM1 Unit := 25 | for _ in [:1000000] do 26 | let _ ← runMetaAsRuleM1 (inferType (.const ``Nat [])) 27 | 28 | def runRuleM1 (m : RuleM1 α) : TacticM α := do 29 | let (ret, state) ← m.run {lctx := ← getLCtx, mctx := ← getMCtx} 30 | setMCtx state.mctx 31 | return ret 32 | 33 | structure State2 where 34 | extra : Nat := 0 35 | 36 | abbrev RuleM2 := StateRefT State2 MetaM 37 | 38 | def ruleM2InferType : RuleM2 Unit := 39 | for _ in [:1000000] do 40 | let _ ← inferType (.const ``Nat []) 41 | 42 | def runRuleM2 (m : RuleM2 α) : TacticM α := do 43 | let (ret, _) ← m.run {} 44 | return ret 45 | 46 | syntax (name := ruleM1Test) "ruleMTest" num : tactic 47 | 48 | @[tactic ruleM1Test] 49 | def runM1Test : Tactic 50 | |`(tactic| ruleMTest $n) => 51 | if n.getNat == 1 then 52 | runRuleM1 ruleM1InferType 53 | else 54 | runRuleM2 ruleM2InferType 55 | | _ => throwError "Unsupported Syntax" 56 | 57 | def foo : Nat := by 58 | ruleMTest 1; -- Change to ```ruleMTest 1``` to see difference 59 | exact 1 60 | 61 | end MonadStackPerformance -------------------------------------------------------------------------------- /Duper/Tests/temp.lean: -------------------------------------------------------------------------------- 1 | import Duper 2 | import Duper.TPTP 3 | 4 | set_option auto.tptp true 5 | set_option auto.tptp.premiseSelection true 6 | set_option auto.tptp.solver.name "zipperposition" 7 | set_option trace.auto.tptp.result true 8 | set_option trace.auto.tptp.printQuery true 9 | set_option auto.native true 10 | 11 | open Lean Auto 12 | 13 | @[rebind Auto.Native.solverFunc] 14 | def Auto.duperPort (lemmas inhLemmas : Array Lemma) : MetaM Expr := do 15 | let (formulas, extraFormulas) ← Duper.autoLemmasToFormulas lemmas 16 | let formulas := formulas.map (fun f => (f.1, f.2.1, f.2.2.1, f.2.2.2, none)) 17 | let extraFormulas := extraFormulas.map (fun f => (f.1, f.2.1, f.2.2.1, f.2.2.2, none)) 18 | Duper.runDuperPortfolioMode formulas extraFormulas .none 19 | { portfolioMode := true, 20 | portfolioInstance := none, 21 | inhabitationReasoning := none, 22 | preprocessing := Duper.PreprocessingOption.NoPreprocessing, -- It would be redundant to enable Auto's preprocessing for Auto calls 23 | includeExpensiveRules := none, 24 | selFunction := none 25 | } 26 | .none 27 | 28 | -- Misc Duper issue: 29 | example (hFin : ∀ x : Nat, ∀ y : Nat, x < y ↔ ∃ a : Fin x, a.1 = x) 30 | (lt_succ : ∀ n : Nat, n < n.succ) (lt_trans : ∀ x y z : Nat, x < y → y < z → x < z) : 31 | ∃ a : Fin (Nat.succ (Nat.succ 0)), a.1 = 0 := by 32 | sorry -- duper [*] {portfolioInstance := 7} -- This causes a stack overflow 33 | 34 | inductive myProd (t1 t2 : Type _) 35 | | mk : t1 → t2 → myProd t1 t2 36 | 37 | open myProd 38 | 39 | -- `unfoldProj` bug in lean-auto 40 | example (sum : myProd Int Int → Int) 41 | (hSum : ∀ x : Int, ∀ y : Int, sum (mk x y) = x + y) 42 | (x : myProd Int Int) : ∃ y : myProd Int Int, sum x < sum y := by 43 | have test : 44 | let _mk_sel0 : myProd Int Int → Int := myProd.rec (motive := fun _ => Int) fun arg0 arg1 => arg0 45 | _mk_sel0 x = 0 := sorry -- This triggers `unfoldProj :: myProd is not a structure` 46 | duper [*] {portfolioInstance := 1} 47 | -------------------------------------------------------------------------------- /Duper/Tests/temp2.lean: -------------------------------------------------------------------------------- 1 | import Duper 2 | import Duper.TPTP 3 | 4 | set_option auto.tptp true 5 | set_option auto.tptp.premiseSelection true 6 | set_option auto.tptp.solver.name "zipperposition" 7 | set_option trace.auto.tptp.result true 8 | set_option trace.auto.tptp.printQuery true 9 | set_option auto.native true 10 | 11 | open Lean Auto 12 | 13 | @[rebind Auto.Native.solverFunc] 14 | def Auto.duperPort (lemmas inhLemmas : Array Lemma) : MetaM Expr := do 15 | let formulas ← Duper.autoLemmasToFormulas lemmas 16 | let formulas := formulas.map (fun f => (f.1, f.2.1, f.2.2.1, f.2.2.2, none)) 17 | Duper.runDuperPortfolioMode formulas .none 18 | { portfolioMode := true, 19 | portfolioInstance := none, 20 | inhabitationReasoning := none, 21 | preprocessing := Duper.PreprocessingOption.NoPreprocessing, -- It would be redundant to enable Auto's preprocessing for Auto calls 22 | includeExpensiveRules := none, 23 | selFunction := none 24 | } 25 | .none 26 | 27 | inductive myType 28 | | const1 : myType 29 | | const2 : myType 30 | 31 | inductive myType2 (t : Type _) 32 | | const3 : t → myType2 t 33 | | const4 : t → myType2 t 34 | 35 | inductive myType3 36 | | const5 : myType3 37 | | const6 : myType3 → myType3 38 | 39 | inductive myType4 (t1 t2 : Type _) 40 | | const7 : t1 → t2 → myType4 t1 t2 41 | 42 | inductive myEmpty (t1 : Type _) (t2 : Type _) 43 | 44 | open myType myType2 myType3 myType4 45 | 46 | set_option duper.collectDatatypes true 47 | set_option trace.duper.rule.datatypeExhaustiveness true 48 | set_option trace.duper.printProof true 49 | set_option trace.duper.collectDatatypes.debug true 50 | 51 | example : ∀ xs : List Unit, ∀ x : Unit, x :: xs ≠ xs := by 52 | -- duper [List.all_nil] {portfolioInstance := 7} 53 | sorry -- Works when datatypeAcyclicity is enabled 54 | 55 | -------------------------------------------------------------------- 56 | /- This example times out when `duper.collectDatatypes` is enabled, look into whether this is a bug or to be expected 57 | 58 | example (f : Nat → Nat → Nat → Nat → Nat → Nat → Nat → Nat) 59 | (g : Nat → Nat → Nat → Nat → Nat → Nat) 60 | (inertia : TPTP.iota → TPTP.iota → TPTP.iota → Prop) 61 | (goodInertia goodChance : TPTP.iota → Prop) 62 | (organization: TPTP.iota → TPTP.iota → Prop) 63 | -- good inertia implies good survival chance 64 | (dummy: ∀ (Y T2 I2 P2 : TPTP.iota), organization Y T2 ∧ inertia Y I2 T2 ∧ goodInertia I2 → goodChance P2) 65 | : ∃ x : Nat, 66 | f (g x x x x x) (g x 1 x x x) (g x x x x x) (g x 1 x x x) (g x x x x x) (g x 1 x x x) (g x x x x x) 67 | = f (g 1 x x x x) (g x x x x x) (g x 1 x x x) (g x x x x x) (g x 1 x x x) (g x x x x x) (g x 1 x x x) := 68 | by duper {portfolioInstance := 7} 69 | -/ 70 | 71 | -------------------------------------------------------------------- 72 | -- Issues relating to Bool support 73 | 74 | set_option trace.duper.saturate.debug true in 75 | theorem Bool.not_eq_true2 (b : Bool) : (¬(b = true)) = (b = false) := by 76 | duper {portfolioInstance := 1} 77 | -- Final Active Set: [∀ (a a_1 : Bool), a = a_1] 78 | 79 | set_option trace.duper.saturate.debug true in 80 | theorem Bool.not_eq_false2 (b : Bool) : (¬(b = false)) = (b = true) := by 81 | duper {portfolioInstance := 1} 82 | -- Final Active Set: [∀ (a a_1 : Bool), a = a_1] 83 | 84 | -------------------------------------------------------------------- 85 | -- Issues relating to type inhabitation reasoning 86 | 87 | set_option trace.duper.timeout.debug true in 88 | theorem forall_comm2 {p : α → β → Prop} : (∀ a b, p a b) ↔ (∀ b a, p a b) := 89 | by sorry -- duper {inhabitationReasoning := true} 90 | 91 | set_option trace.duper.timeout.debug true in 92 | theorem exists_comm2 {p : α → β → Prop} : (∃ a b, p a b) ↔ (∃ b a, p a b) := 93 | by sorry -- duper {inhabitationReasoning := true} 94 | 95 | set_option trace.duper.timeout.debug true in 96 | theorem forall_apply_eq_imp_iff2 {f : α → β} {p : β → Prop} : 97 | (∀ b a, f a = b → p b) ↔ ∀ a, p (f a) := by sorry -- duper {inhabitationReasoning := true} 98 | 99 | theorem exists_exists_and_eq_and {f : α → β} {p : α → Prop} {q : β → Prop} : 100 | ∃ a, p a ∧ q (f a) → (∃ b, (∃ a, p a ∧ f a = b) ∧ q b) := by duper 101 | 102 | theorem exists_exists_eq_and {f : α → β} {p : β → Prop} : 103 | ∃ a, p (f a) → (∃ b, (∃ a, f a = b) ∧ p b) := by duper 104 | 105 | -------------------------------------------------------------------- 106 | theorem forall₂_true_iff1 {β : α → Sort} : (∀ a, β a → True) ↔ True := by duper [*] 107 | 108 | -- Look into why Duper can solve `forall₂_true_iff1` but not `forall₂_true_iff2` 109 | theorem forall₂_true_iff2 {β : α → Sort _} : (∀ a, β a → True) ↔ True := by duper [*] 110 | 111 | -------------------------------------------------------------------- 112 | set_option trace.duper.saturate.debug true in 113 | set_option duper.throwPortfolioErrors true in 114 | theorem exists₂_comm 115 | {ι₁ ι₂ : Sort} {κ₁ : ι₁ → Sort} {κ₂ : ι₂ → Sort} {p : ∀ i₁, κ₁ i₁ → ∀ i₂, κ₂ i₂ → Prop} : 116 | (∃ i₁ j₁ i₂ j₂, p i₁ j₁ i₂ j₂) ↔ ∃ i₂ j₂ i₁ j₁, p i₁ j₁ i₂ j₂ := by 117 | duper [*] {preprocessing := no_preprocessing} 118 | -------------------------------------------------------------------------------- /Duper/Tests/test_Order.lean: -------------------------------------------------------------------------------- 1 | import Duper.Order 2 | 3 | 4 | namespace Duper.Order.Test 5 | open Lean 6 | 7 | opaque a : Nat 8 | opaque b : Nat 9 | opaque c : Nat 10 | opaque f : Nat → Nat 11 | opaque g : Nat → Nat → Nat 12 | 13 | def consts := #[mkConst ``a, mkConst ``b, mkConst ``c] 14 | 15 | partial def constructRandomExpr (fuel : Nat) : MetaM (Expr × Nat) := do 16 | let fuel := fuel - 1 17 | let i := if fuel == 0 then 0 else ← IO.rand 0 2 18 | if i == 0 then 19 | return (consts[← IO.rand 0 (consts.size - 1)]!, fuel) 20 | else if i == 1 then 21 | let (a, fuel) ← constructRandomExpr fuel 22 | return (mkApp (mkConst ``f) a, fuel) 23 | else 24 | let (a, fuel) ← constructRandomExpr fuel 25 | let (b, fuel) ← constructRandomExpr fuel 26 | return (mkApp2 (mkConst ``g) a b, fuel) 27 | 28 | 29 | #eval show Elab.TermElabM _ from do 30 | for _ in [:10] do 31 | let (a, _) ← constructRandomExpr (fuel := 20) 32 | let (b, _) ← constructRandomExpr (fuel := 20) 33 | let cab ← kbo a b {} false 34 | let cba ← kbo b a {} false 35 | logInfo m!"{a}\n{cab}\n{b}" 36 | 37 | -- Check symmetry 38 | guard $ cab == cba.flip 39 | 40 | end Duper.Order.Test -------------------------------------------------------------------------------- /Duper/Tests/test_PUZ.lean: -------------------------------------------------------------------------------- 1 | import Duper.TPTP 2 | import Duper.Tactic 3 | 4 | tptp PUZ018_1 "../TPTP-v8.0.0/Problems/PUZ/PUZ018_1.p" 5 | by duper -- Det timeout 6 | 7 | tptp PUZ031_1 "../TPTP-v8.0.0/Problems/PUZ/PUZ031_1.p" 8 | by duper -- Failed to synthesize inhabited_plant 9 | 10 | tptp PUZ031_1_modified "../TPTP-v8.0.0/Problems/PUZ/PUZ031_1.p" by 11 | have inhabited_plant : Inhabited plant := sorry 12 | have inhabited_snail : Inhabited snail := sorry 13 | have inhabited_grain : Inhabited grain := sorry 14 | have inhabited_bird : Inhabited bird := sorry 15 | have inhabited_fox : Inhabited fox := sorry 16 | have Inhabited_wolf : Inhabited wolf := sorry 17 | duper -- If these instances are not provided, duper will fail 18 | -- Previously: Error when reconstructing clausification 19 | 20 | tptp PUZ137_8 "../TPTP-v8.0.0/Problems/PUZ/PUZ137_8.p" 21 | by duper -- Succeeds 22 | 23 | tptp PUZ081_8 "../TPTP-v8.0.0/Problems/PUZ/PUZ081_8.p" 24 | by duper -- Succeeds 25 | 26 | tptp PUZ083_8 "../TPTP-v8.0.0/Problems/PUZ/PUZ083_8.p" 27 | by duper -- Succeeds 28 | 29 | tptp PUZ134_2 "../TPTP-v8.0.0/Problems/PUZ/PUZ134_2.p" 30 | by duper -- Det timeout 31 | 32 | tptp PUZ135_2 "../TPTP-v8.0.0/Problems/PUZ/PUZ135_2.p" 33 | by duper -- Det timeout 34 | 35 | tptp PUZ082_8 "../TPTP-v8.0.0/Problems/PUZ/PUZ082_8.p" 36 | by duper -- Succeeds 37 | 38 | tptp PUZ130_1 "../TPTP-v8.0.0/Problems/PUZ/PUZ130_1.p" 39 | by duper -- Succeeds 40 | 41 | tptp PUZ131_1 "../TPTP-v8.0.0/Problems/PUZ/PUZ131_1.p" 42 | by duper -- Succeeds 43 | 44 | tptp PUZ135_1 "../TPTP-v8.0.0/Problems/PUZ/PUZ135_1.p" 45 | by duper -- Det timeout 46 | 47 | tptp PUZ134_1 "../TPTP-v8.0.0/Problems/PUZ/PUZ134_1.p" 48 | by duper -- Det timeout 49 | 50 | tptp PUZ139_1 "../TPTP-v8.0.0/Problems/PUZ/PUZ139_1.p" 51 | by duper -- Succeeds 52 | 53 | /- 54 | tptp PUZ133_2 "../TPTP-v8.0.0/Problems/PUZ/PUZ133=2.p" -- Parsing issue 55 | by duper PUZ133_2 56 | -/ 57 | 58 | tptp PUZ012_1 "../TPTP-v8.0.0/Problems/PUZ/PUZ012_1.p" 59 | by duper -- See TODO 60 | 61 | #print axioms PUZ018_1 62 | #print axioms PUZ031_1 63 | #print axioms PUZ137_8 64 | --#print axioms PUZ081_8 65 | #print axioms PUZ083_8 66 | #print axioms PUZ134_2 67 | #print axioms PUZ135_2 68 | #print axioms PUZ082_8 69 | #print axioms PUZ130_1 70 | #print axioms PUZ131_1 71 | #print axioms PUZ135_1 72 | #print axioms PUZ134_1 73 | #print axioms PUZ139_1 74 | --#print axioms PUZ133_2 75 | #print axioms PUZ012_1 -------------------------------------------------------------------------------- /Duper/Tests/test_boolSimp.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | import Duper.TPTP 3 | 4 | -- Note: These tests only effectively test boolSimp when identBoolHoist is disabled 5 | set_option trace.duper.rule.boolSimp true 6 | 7 | theorem boolSimpRule1Test (p : Prop) (h : (p ∨ p ∨ p ∨ p) = q) : p = q := 8 | by duper [*] 9 | 10 | theorem boolSimpRule2Test (p q : Prop) (h : (¬p ∨ p) = q) : q := 11 | by duper [*] 12 | 13 | theorem boolSimpRule2SymTest (p q : Prop) (h : (p ∨ ¬p) = q) : q := 14 | by duper [*] 15 | 16 | theorem boolSimpRule3Test (p q : Prop) (h : (p ∨ True) = q) : q := 17 | by duper [*] 18 | 19 | theorem boolSimpRule3SymTest (p q : Prop) (h : (True ∨ p) = q) : q := 20 | by duper [*] 21 | 22 | theorem boolSimpRule4Test (p q : Prop) (h : (p ∨ False) = (q ∨ False)) : p = q := 23 | by duper [*] 24 | 25 | theorem boolSimpRule4SymTest (p q : Prop) (h : (False ∨ p) = (False ∨ q)) : p = q := 26 | by duper [*] 27 | 28 | theorem boolSimpRule5Test (p q : Prop) (h : p = (q = q)) : p := 29 | by duper [*] 30 | 31 | theorem boolSimpRule6Test (p q : Prop) (h : (p = True) = (q = True)) : p = q := 32 | by duper [*] 33 | 34 | theorem boolSimpRule6SymTest (p q : Prop) (h : (True = p) = (True = q)) : p = q := 35 | by duper [*] 36 | 37 | theorem boolSimpRule7Test (p q : Prop) (h : p = Not False) : p := 38 | by duper [*] 39 | 40 | theorem boolSimpRule8Test (p : Prop) (h : (p ∧ p ∧ p ∧ p) = q) : p = q := 41 | by duper [*] 42 | 43 | theorem boolSimpRule9Test (p q : Prop) (h : (¬p ∧ p) = q) : ¬q := 44 | by duper [*] 45 | 46 | theorem boolSimpRule9SymTest (p q : Prop) (h : (p ∧ ¬p) = q) : ¬q := 47 | by duper [*] 48 | 49 | theorem boolSimpRule10Test (p q : Prop) (h : (p ∧ True) = q) : p = q := 50 | by duper [*] 51 | 52 | theorem boolSimpRule10SymTest (p q : Prop) (h : (True ∧ p) = q) : p = q := 53 | by duper [*] 54 | 55 | theorem boolSimpRule11Test (p q : Prop) (h : (p ∧ False) = q) : ¬q := 56 | by duper [*] 57 | 58 | theorem boolSimpRule11SymTest (p q : Prop) (h : (False ∧ p) = q) : ¬q := 59 | by duper [*] 60 | 61 | theorem boolSimpRule12Test (p q : Prop) (h : p = (q ≠ q)) : ¬p := 62 | by duper [*] 63 | 64 | theorem boolSimpRule13Test (p q : Prop) (h : (p = False) = (q = False)) : p = q := 65 | by duper [*] 66 | 67 | theorem boolSimpRule13SymTest (p q : Prop) (h : (False = p) = (False = q)) : p = q := 68 | by duper [*] 69 | 70 | theorem boolSimpRule14Test (p q : Prop) (h : p = Not True) : ¬p := 71 | by duper [*] 72 | 73 | theorem boolSimpRule15Test (p q : Prop) (h : (¬¬p) = q) : p = q := 74 | by duper [*] 75 | 76 | theorem boolSimpRule16Test (p q : Prop) (h : (True → p) = q) : p = q := 77 | by duper [*] 78 | 79 | theorem boolSimpRule17Test (p q : Prop) (h : (False → p) = q) : q := 80 | by duper [*] 81 | 82 | theorem boolSimpRule18Test (p q : Prop) (h : (p → False) = (q → False)) : p = q := 83 | by duper [*] 84 | 85 | theorem boolSimpRule19Test (p q : Prop) (h : (p → True) = q) : q := 86 | by duper [*] 87 | 88 | theorem boolSimpRule19Test2 (α) (q : Prop) (h : (∀ _ : α, True) = q) : q := 89 | by duper [*] 90 | 91 | theorem boolSimpRule20Test (p q : Prop) (h : (p → ¬p) = (q → ¬q)) : p = q := 92 | by duper [*] 93 | 94 | theorem boolSimpRule21Test (p q : Prop) (h : (¬p → p) = (¬q → q)) : p = q := 95 | by duper [*] 96 | 97 | theorem boolSimpRule22Test (p q : Prop) (h : (p → p) = q) : q := 98 | by duper [*] 99 | 100 | theorem boolSimpRule23Test (f : Prop → Prop) (q : Prop) (hq : q) (h : (∀ p : Prop, f p) = q) : f True := 101 | by duper [*] 102 | 103 | theorem boolSimpRule24Test (f : Prop → Prop) (q : Prop) (hq : q) (h : (∃ p : Prop, f p) = q) : (f True) ∨ (f False) := 104 | by duper [*] 105 | 106 | theorem boolSimpRule25Test (p q r : Prop) (h : (p → ¬q → q → p → False) = r) : r := 107 | by duper [*] 108 | 109 | theorem boolSimpRule26Test (a b c shared x y z r : Prop) (h : (a → b → shared → c → (x ∨ shared ∨ y ∨ z)) = r) : r := 110 | by duper [*] 111 | 112 | theorem boolSimpRule27Test (a b c shared x y z r : Prop) (h : ((a ∧ b ∧ shared ∧ c) → (x ∨ shared ∨ y ∨ z)) = r) : r := 113 | by duper [*] 114 | 115 | theorem boolSimpRule28Test (p q r : Prop) (h : (p ↔ r) = (q ↔ r)) : p = q := 116 | by duper [*] 117 | -------------------------------------------------------------------------------- /Duper/Tests/test_continuity.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | 3 | axiom Real : Type 4 | axiom dist : Real → Real → Real 5 | axiom add : Real → Real → Real 6 | axiom lt : Real → Real → Prop 7 | axiom zero : Real 8 | axiom one : Real 9 | 10 | def continuous_at (f : Real → Real) (x₀ : Real) := 11 | ∀ ε : Real, lt zero ε → ∃ δ : Real, lt zero δ ∧ ∀ x, lt (dist x x₀) δ → lt (dist (f x) (f x₀)) ε 12 | 13 | def continuous (f : Real → Real) := ∀ x, continuous_at f x 14 | 15 | def uniformly_continuous (f : Real → Real) := 16 | ∀ ε : Real, lt zero ε → ∃ δ : Real, lt zero δ ∧ ∀ x y, lt (dist x y) δ → lt (dist (f x) (f y)) ε 17 | 18 | theorem dist_self (a : Real) : dist a a = zero := sorry 19 | 20 | theorem zero_lt_one (a : Real) : lt zero one := sorry 21 | 22 | example (a : Real) : continuous (λ _ => a) := 23 | by simp only [continuous_at, continuous]; duper [dist_self] 24 | 25 | example (a : Real) : continuous (λ _ => a) := 26 | by duper [continuous_at, continuous, dist_self] 27 | 28 | example (a : Real) : uniformly_continuous (λ _ => a) := 29 | by simp only [uniformly_continuous]; duper [dist_self, zero_lt_one] 30 | 31 | example (a : Real) : uniformly_continuous (λ _ => a) := 32 | by duper [uniformly_continuous, dist_self, zero_lt_one] 33 | 34 | example (a : Real) : continuous (λ x => x) := 35 | by simp only [continuous_at, continuous]; duper [dist_self] 36 | 37 | example (a : Real) : continuous (λ x => x) := 38 | by duper [continuous_at, continuous, dist_self] 39 | 40 | example (a : Real) : uniformly_continuous (λ x => x) := 41 | by duper [uniformly_continuous, dist_self, zero_lt_one] 42 | 43 | example (f : Real → Real) : uniformly_continuous f → continuous f := 44 | by simp only [continuous, continuous_at, uniformly_continuous]; duper 45 | 46 | example (f : Real → Real) : uniformly_continuous f → continuous f := 47 | by duper [continuous, continuous_at, uniformly_continuous] 48 | -------------------------------------------------------------------------------- /Duper/Tests/test_higher_order.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | import Duper.TPTP 3 | 4 | set_option maxHeartbeats 20000 5 | tptp NUM020_1 "../TPTP-v8.0.0/Problems/NUM/NUM020^1.p" 6 | by duper [*] 7 | 8 | tptp AGT033 "../TPTP-v8.0.0/Problems/AGT/AGT033^1.p" by sorry 9 | 10 | set_option trace.Meta.debug true in 11 | tptp ANA047 "../TPTP-v8.0.0/Problems/ANA/ANA047^1.p" by sorry 12 | 13 | tptp DAT113 "../TPTP-v8.0.0/Problems/DAT/DAT113^1.p" by sorry 14 | 15 | -- Higher order tests 16 | example 17 | (three six : (Nat → Nat) → Nat → Nat) 18 | (succ : ((Nat → Nat) → Nat → Nat) → ((Nat → Nat) → Nat → Nat)) 19 | (hsucc_ax: succ = fun N X Y => X (N X Y)) 20 | (plus mult : ((Nat → Nat) → Nat → Nat) → ((Nat → Nat) → Nat → Nat) → ((Nat → Nat) → Nat → Nat)) 21 | (hmult_ax: mult = fun M N X Y => M (N X) Y) 22 | (hthree_ax: three = fun X Y => X (X (X Y))) 23 | (hthm: ¬∃ N, mult N three = three) : False := by duper [*] 24 | 25 | -- Ex27 is example 27 from https://matryoshka-project.github.io/pubs/hosup_report.pdf 26 | set_option trace.duper.printProof true in 27 | set_option trace.duper.rule.neHoist true in 28 | set_option trace.duper.proofReconstruction true in 29 | theorem ex27 (t : Type) (g : Prop → t) (h : t → t) (A B : t) 30 | (eq1 : A ≠ B) 31 | (eq2 : ∀ y : t → t, h (y B) ≠ h (g False) ∨ h (y A) ≠ h (g True)) : False := 32 | by duper [*] 33 | -------------------------------------------------------------------------------- /Duper/Tests/test_inhabitationReasoning.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | import Duper.TPTP 3 | 4 | set_option inhabitationReasoning true 5 | set_option trace.duper.typeInhabitationReasoning.debug true 6 | 7 | theorem optionTest1 (t : Type) (f : Option t) : ∃ x : Option t, True := by duper 8 | 9 | theorem optionTest2 : ∀ t : Type, ∃ x : Option t, True := by duper 10 | 11 | theorem nonemptyHypTest (t : Type) (eh : Nonempty t = True ∧ True) (h : ∀ x : t, False ≠ False) : False := by 12 | duper [eh, h] 13 | 14 | -- Needs to synthesize Inhabited (Fin x) 15 | theorem finTest (x : Nat) (f : Fin x → Fin x) 16 | (h : ∃ y : Fin x, ∀ z : Fin x, (f y ≠ y) ∧ (f z = y)) : False := by duper [h] 17 | 18 | -- Needs to synthesize Inhabited (Fin default) 19 | set_option trace.duper.timeout.debug true in 20 | set_option maxHeartbeats 5000 in 21 | theorem finTest2 22 | (h : ∀ x : Nat, ∃ f : Fin x → Fin x, ∃ y : Fin x, ∀ z : Fin x, (f y ≠ y) ∧ (f z = y)) : False := by 23 | duper [h] {portfolioInstance := 7} 24 | 25 | -- Needs to synthesize Inhabited (Fin [anonymous]) (pretty sure [anonymous] is a skolem var) 26 | set_option trace.duper.saturate.debug true in 27 | theorem finTest3 (mult_Nats : ∃ y : Nat, y ≠ 0) 28 | (h : ∀ x : Nat, x ≠ 0 → ∃ f : Fin x → Fin x, ∃ y : Fin x, ∀ z : Fin x, (f y ≠ y) ∧ (f z = y)) : False := by 29 | duper [h] 30 | 31 | -- Needs to synthesize Inhabited person 32 | set_option trace.duper.proofReconstruction true in 33 | theorem barber_paradox1 {person : Type} {shaves : person → person → Prop} 34 | (h : ∃ b : person, ∀ p : person, (shaves b p ↔ (¬ shaves p p))) : False := 35 | by duper [h] 36 | 37 | theorem letDecBug {t : Type} (h : (∀ p : t, p = p) = False) : False := 38 | by duper [h] 39 | 40 | -- Interesting type inhabited examples (they require more advanced reasoning about type inhabitation) 41 | example : ((∃ (A B : Type) (f : B → A) (x : B), f x = f x) = True) := 42 | by duper {portfolioInstance := 7} 43 | 44 | example : ((∃ (A B : Type) (f : B → A) (x : B), f x = f x) = True) := 45 | by duper {portfolioInstance := 8} 46 | 47 | example : ∃ (A : Type) (B : A → Type) (f : ∀ (a : A), B a) (x : A), (f x = f x) = True := 48 | by duper {portfolioInstance := 7} 49 | 50 | example : ∃ (A : Type) (B : A → Type) (f : ∀ (a : A), B a) (x : A), (f x = f x) = True := 51 | by duper {portfolioInstance := 8} 52 | 53 | set_option trace.duper.proofReconstruction true in 54 | example : ((∀ (A : Type) (f : Nat → A) (x : Nat), f x = f x) = True) := 55 | by duper 56 | -------------------------------------------------------------------------------- /Duper/Unif.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | 3 | set_option linter.unusedVariables false 4 | 5 | open Lean 6 | open Lean.Meta 7 | 8 | partial def Lean.Meta.fastUnify (l : Array (Expr × Expr)) : MetaM Bool := do 9 | Core.checkMaxHeartbeats "unify" 10 | let state ← saveState 11 | try 12 | for (t, s) in l do 13 | let t_type := (← instantiateMVars (← inferType t)) 14 | let s_type := (← instantiateMVars (← inferType s)) 15 | if ¬ (← unify1 t_type s_type) then 16 | state.restore 17 | return false 18 | else if ← unify1 t s then 19 | continue 20 | else 21 | state.restore 22 | return false 23 | return true 24 | catch ex => 25 | state.restore 26 | throw ex 27 | where 28 | unify1 (s t : Expr) : MetaM Bool := do 29 | let s ← instantiateMVars s --TODO: instantiate more lazily? 30 | let t ← instantiateMVars t 31 | if s == t then return true else 32 | s.withApp fun f ss => t.withApp fun g tt => 33 | match f, g with 34 | | Expr.fvar fid, Expr.fvar gid => 35 | if fid == gid then 36 | unifyArgs ss tt 37 | else 38 | return false 39 | | Expr.const fname flvls, Expr.const gname glvls => do 40 | if fname == gname && 41 | (← (flvls.zip glvls).allM (fun (flvl, glvl) => isLevelDefEq flvl glvl)) then 42 | unifyArgs ss tt 43 | else 44 | return false 45 | | Expr.mvar mVarId .., Expr.fvar .. => 46 | unifyFlexRigid s t 47 | | Expr.mvar mVarId .., Expr.const .. => 48 | unifyFlexRigid s t 49 | | Expr.fvar .., Expr.mvar mVarId .. => 50 | unifyFlexRigid t s 51 | | Expr.const .., Expr.mvar mVarId .. => 52 | unifyFlexRigid t s 53 | | Expr.mvar mVarId .., Expr.mvar .. => 54 | unifyFlexFlex s t 55 | -- Adding cases involving forallE (treating it similar to Expr.const) 56 | | Expr.forallE .., Expr.forallE .. => 57 | return s == t 58 | | Expr.mvar .., Expr.forallE .. => 59 | unifyFlexRigid s t 60 | | Expr.forallE .., Expr.mvar .. => 61 | unifyFlexRigid t s 62 | | _, _ => return false 63 | unifyArgs (ss tt : Array Expr) : MetaM Bool := do 64 | if tt.size == ss.size then 65 | fastUnify (tt.zip ss) 66 | else 67 | return false 68 | unifyFlexFlex (s t : Expr) : MetaM Bool := do 69 | s.withApp fun f ss => t.withApp fun g tt => do 70 | if tt.isEmpty && ss.isEmpty 71 | then 72 | if f == g 73 | then return true 74 | else 75 | Lean.MVarId.assign f.mvarId! t 76 | return true 77 | else return false 78 | unifyFlexRigid (s t : Expr) : MetaM Bool := do 79 | s.withApp fun f ss => t.withApp fun g tt => do 80 | match ← getExprMVarAssignment? f.mvarId! with 81 | | some e => 82 | unify1 (mkAppN e ss) t 83 | | none => do 84 | if ss.isEmpty 85 | then 86 | if ← occursCheck f.mvarId! t 87 | then 88 | Lean.MVarId.assign f.mvarId! t 89 | return true 90 | else 91 | return false 92 | else return false 93 | -------------------------------------------------------------------------------- /Duper/Util/Debugging.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | open Lean 3 | 4 | namespace Duper 5 | 6 | -- Check heartbeat with given maxheartbeat number 7 | 8 | def checkGivenHeartbeats (moduleName : String) (max : Nat) : CoreM Unit := do 9 | Core.checkMaxHeartbeatsCore moduleName `maxHeartbeats (max * 1000) 10 | 11 | def checkGivenMicroHeartbeats (moduleName : String) (max : Nat) : CoreM Unit := do 12 | Core.checkMaxHeartbeatsCore moduleName `maxHeartbeats max 13 | 14 | def getNormalizedMicroHeartbeats : CoreM Nat := do 15 | let num ← IO.getNumHeartbeats 16 | let init := (← read).initHeartbeats 17 | return num - init 18 | 19 | -- Options that can be used in debugging 20 | 21 | register_option duperDebugOpt1 : Nat := { 22 | defValue := 0 23 | descr := "Duper debugging option 1" 24 | } 25 | 26 | def getDuperDebugOpt1 : CoreM Nat := do 27 | let opts ← getOptions 28 | return duperDebugOpt1.get opts 29 | 30 | register_option duperDebugOpt2 : Nat := { 31 | defValue := 0 32 | descr := "Duper debugging option 2" 33 | } 34 | 35 | def getDuperDebugOpt2 : CoreM Nat := do 36 | let opts ← getOptions 37 | return duperDebugOpt2.get opts -------------------------------------------------------------------------------- /Duper/Util/DeeplyOccurringVars.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import Duper.Util.Misc 3 | import Duper.Order 4 | import Duper.MClause 5 | import Duper.Expr 6 | 7 | set_option linter.unusedVariables false 8 | 9 | namespace Duper 10 | open Lean 11 | 12 | /-- Note: This function currently excludes lambdas that appear in forallE, lam, or letE types. -/ 13 | def occursInLambdaNotDirectlyBelowQuantifier (c : MClause) (t : Expr) : Bool := 14 | -- For the visit function, inLambda is only set to true if the lambda visit is currently in is not directly below a quantifier 15 | let rec @[specialize] visit (e : Expr) (directlyBelowQuantifier : Bool) (inLambda : Bool) := 16 | if inLambda && e == t then true else 17 | match e with 18 | | Expr.forallE _ ty b _ => visit b true inLambda 19 | | Expr.lam _ ty b _ => 20 | if directlyBelowQuantifier then visit b false false 21 | else visit b false true 22 | | Expr.mdata _ e => visit e directlyBelowQuantifier inLambda 23 | | Expr.letE _ t v b _ => 24 | -- Not sure if this is the best way to handle letE expressions 25 | -- For example, in `∃ x : ty, let y := s; t`, `t` is sort of directly below the quantifier even though it's technically under the letE expression 26 | visit v false inLambda || visit b false inLambda 27 | | Expr.app f a => visit f false inLambda || visit a false inLambda 28 | | Expr.proj _ _ e => visit e false inLambda 29 | | _ => false 30 | c.fold (fun b e => b || visit e false false) false 31 | 32 | def occursInArgumentOfAppliedVariable (c : MClause) (t : Expr) : Bool := 33 | let rec @[specialize] visit (e : Expr) (inArgumentOfAppliedVariable : Bool) := 34 | if inArgumentOfAppliedVariable && e == t then true else 35 | match e with 36 | | Expr.forallE _ d b _ => visit d inArgumentOfAppliedVariable || visit b inArgumentOfAppliedVariable 37 | | Expr.lam _ d b _ => visit d inArgumentOfAppliedVariable || visit b inArgumentOfAppliedVariable 38 | | Expr.mdata _ e => visit e inArgumentOfAppliedVariable 39 | | Expr.letE _ t v b _ => visit t inArgumentOfAppliedVariable || visit v inArgumentOfAppliedVariable || visit b inArgumentOfAppliedVariable 40 | | Expr.app f a => 41 | if visit f inArgumentOfAppliedVariable then true 42 | else if f.getAppFn.isMVar' then visit a true 43 | else visit a inArgumentOfAppliedVariable 44 | | Expr.proj _ _ e => visit e inArgumentOfAppliedVariable 45 | | _ => false 46 | c.fold (fun b e => b || visit e false) false 47 | 48 | /-- Returns true if `t` is an mvar that occurs deeply in MClause `c`. 49 | 50 | A variable occurs deeply in an mclause `c` if it occurs inside an argument of an applied 51 | variable or inside a λ-expression that is not directly below a quantifier. -/ 52 | def isDeep (c : MClause) (t : Expr) : Bool := 53 | t.isMVar' && (occursInArgumentOfAppliedVariable c t || occursInLambdaNotDirectlyBelowQuantifier c t) 54 | 55 | /-- A syntactic overapproximation of fluid or deep terms -/ 56 | def isFluidOrDeep (c : MClause) (t : Expr) : Bool := Order.isFluid t || isDeep c t 57 | -------------------------------------------------------------------------------- /Duper/Util/MessageData.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | open Lean 3 | 4 | namespace Duper 5 | 6 | def ExprPairToMessageData : Expr × Expr → MessageData 7 | | (e1, e2) => MessageData.compose m!"({e1}, " m!"{e2})" 8 | 9 | def ListToMessageData (as : List α) (f : α → MessageData) : MessageData := 10 | .compose m!"[" (.compose (go as) m!"]") 11 | where go as := 12 | match as with 13 | | .nil => m!"" 14 | | .cons a .nil => f a 15 | | .cons a as@(.cons _ _) => .compose (f a) (.compose ", " (go as)) 16 | 17 | def ArrayToMessageData (as : Array α) (f : α → MessageData) : MessageData := 18 | ListToMessageData as.toList f 19 | -------------------------------------------------------------------------------- /Duper/Util/OccursCheck.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Duper development team. All rights reserved. 3 | Author: Duper development team 4 | -/ 5 | import Lean.MetavarContext 6 | import Lean 7 | 8 | namespace Lean 9 | 10 | /-- 11 | Return true if `e` **must not** contain `mvarId` directly or indirectly 12 | This function considers assigments and delayed assignments. -/ 13 | partial def mustNotOccursCheck (mvarId : MVarId) (e : Expr) : MetaM Bool := do 14 | if !e.hasExprMVar then 15 | return true 16 | else 17 | match (← visit e |>.run |>.run {}) with 18 | | (.ok .., _) => return true 19 | | (.error .., _) => return false 20 | where 21 | visitMVar (mvarId' : MVarId) : ExceptT Unit (StateT ExprSet MetaM) Unit := do 22 | if mvarId == mvarId' then 23 | throw (self:=instMonadExceptOfMonadExceptOf _ _) () -- found 24 | else 25 | let mvarTy' ← Meta.inferType (mkMVar mvarId') 26 | -- Modification : We also need to check whether metavariables 27 | -- might depend on the metavariable being checked, so we also 28 | -- visit the type of the metavariable declaration 29 | visit mvarTy' 30 | match (← getExprMVarAssignment? mvarId') with 31 | | some v => visit v 32 | | none => 33 | match (← getDelayedMVarAssignment? mvarId') with 34 | | some d => visitMVar d.mvarIdPending 35 | | none => return () 36 | 37 | visit (e : Expr) : ExceptT Unit (StateT ExprSet MetaM) Unit := do 38 | if !e.hasExprMVar then 39 | return () 40 | else if (← get).contains e then 41 | return () 42 | else 43 | modify fun s => s.insert e 44 | match e with 45 | | Expr.proj _ _ s => visit s 46 | | Expr.forallE _ d b _ => visit d; visit b 47 | | Expr.lam _ d b _ => visit d; visit b 48 | | Expr.letE _ t v b _ => visit t; visit v; visit b 49 | | Expr.mdata _ b => visit b 50 | | Expr.app f a => visit f; visit a 51 | | Expr.mvar mvarId => visitMVar mvarId 52 | | _ => return () 53 | 54 | private def Expr.isOccursRigid [Monad m] [MonadLCtx m] : Expr → m Bool 55 | | .bvar _ => pure true 56 | -- fvar : rigid if it's `cdecl`, not rigid if it's `ldecl` 57 | | e@(.fvar id) => do 58 | let lctx ← getLCtx 59 | match lctx.find? id with 60 | | none => panic! s!"Expr.isOccursRigid : Free variable {e} is not in local context" 61 | | some (.cdecl _ _ _ _ _ _) => pure true 62 | | some (.ldecl _ _ _ _ _ _ _) => pure false 63 | | .mvar _ => pure false 64 | | .sort _ => pure true 65 | | .const _ _ => pure true 66 | | .app _ _ => pure false 67 | | .lam _ _ _ _ => pure false 68 | | .forallE _ _ _ _ => pure true 69 | | .letE _ _ _ _ _ => pure false 70 | | .lit _ => pure true 71 | | .mdata _ e => e.isOccursRigid 72 | | .proj _ _ _ => pure false 73 | 74 | /-- 75 | Return true if `e` **must** contain `mvarId` as a subterm directly or 76 | indirectly after further mvar instantiation and β-reduction 77 | This function considers assigments and delayed assignments. -/ 78 | partial def mustOccursCheck [Monad m] [MonadMCtx m] [MonadLCtx m] (mvarId : MVarId) (e : Expr) : m Bool := do 79 | if !e.hasExprMVar then 80 | return false 81 | else 82 | match (← visit e |>.run |>.run {}) with 83 | | (.ok .., _) => return false 84 | | (.error .., _) => return true 85 | where 86 | visitMVar (mvarId' : MVarId) : ExceptT Unit (StateT ExprSet m) Unit := do 87 | if mvarId == mvarId' then 88 | throw () -- found 89 | else 90 | match (← getExprMVarAssignment? mvarId') with 91 | | some v => visit v 92 | | none => 93 | match (← getDelayedMVarAssignment? mvarId') with 94 | | some d => visitMVar d.mvarIdPending 95 | | none => return () 96 | 97 | visit (e : Expr) : ExceptT Unit (StateT ExprSet m) Unit := do 98 | if !e.hasExprMVar then 99 | return () 100 | else if (← get).contains e then 101 | return () 102 | else 103 | modify fun s => s.insert e 104 | match e with 105 | | Expr.proj _ _ s => visit s 106 | | Expr.forallE _ d b _ => visit d; visit b 107 | -- does not visit type of lambda binders 108 | | Expr.lam _ _ b _ => visit b 109 | -- does not visit type and value of "let" binders 110 | | Expr.letE _ _ _ b _ => visit b 111 | | Expr.mdata _ b => visit b 112 | | Expr.app _ _ => 113 | let f := e.getAppFn 114 | visit f 115 | if ← f.isOccursRigid then 116 | let args := e.getAppArgs 117 | for arg in args do 118 | visit arg 119 | | Expr.mvar mvarId => visitMVar mvarId 120 | | _ => return () 121 | 122 | end Lean 123 | -------------------------------------------------------------------------------- /Duper/Util/Reduction.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | 3 | namespace Duper 4 | open Std 5 | open Lean 6 | open Meta 7 | open Core 8 | 9 | initialize 10 | registerTraceClass `duper.preprocessing.debug 11 | 12 | /-- Note: This option should never be enabled if Duper will later call the monomorphization procedure 13 | because attempting to reduce typeclass instances can interfere with monomorphization. It also does 14 | not appear to provide observable benefit even when the monomorphization procedure is not being called, 15 | so it is currently disabled at all times. The option is just left around (rather than deleted entirely) 16 | for testing purposes. -/ 17 | register_option reduceInstances : Bool := { 18 | defValue := false 19 | descr := "Whether to eliminate mdata and apply whnf to instances" 20 | } 21 | 22 | def getReduceInstances (opts : Options) : Bool := 23 | reduceInstances.get opts 24 | 25 | def getReduceInstancesM : CoreM Bool := do 26 | let opts ← getOptions 27 | return getReduceInstances opts 28 | 29 | /-- This function is expensive and should only be used in preprocessing -/ 30 | partial def preprocessFact (fact : Expr) : MetaM Expr := do 31 | if (← getReduceInstancesM) then 32 | let red (e : Expr) : MetaM TransformStep := do 33 | let e := e.consumeMData 34 | let e ← whnf e 35 | return .continue e 36 | -- Reduce 37 | trace[duper.preprocessing.debug] "fact before preprocessing: {fact}" 38 | let fact ← withTransparency .instances <| Meta.transform fact (pre := red) (usedLetOnly := false) 39 | let restoreNE (e : Expr) : MetaM TransformStep := do 40 | match e with 41 | | .app (.const ``Not []) (.app (.app (.app (.const ``Eq lvls) ty) e₁) e₂) => 42 | return .visit (.app (.app (.app (.const ``Ne lvls) ty) e₁) e₂) 43 | | _ => return .continue 44 | -- Restore ≠, i.e., ¬ a = b ⇒ a ≠ b 45 | -- If we don't do this, it seems that clausification will become more inefficient 46 | let fact ← Core.transform fact (pre := restoreNE) 47 | let fact ← zetaReduce fact 48 | trace[duper.preprocessing.debug] "fact after preprocessing: {fact}" 49 | return fact 50 | else 51 | trace[duper.preprocessing.debug] "reduceInstances option is set to false, only applying zeta reduction" 52 | zetaReduce fact 53 | 54 | /-- Eta-expand a beta-reduced expression. This function is currently unused -/ 55 | partial def etaLong (e : Expr) : MetaM Expr := do 56 | match e with 57 | | Expr.forallE .. => 58 | forallTelescope e fun xs b => do 59 | let bl ← etaLong b 60 | mkForallFVarsEtaLong xs bl 61 | | Expr.lam .. => 62 | lambdaTelescope e fun xs b => do 63 | let bl ← etaLong b 64 | mkLambdaFVarsEtaLong xs bl 65 | | Expr.letE .. => do 66 | throwError "There should be no let binders because they should have been reduced during preprocessing" 67 | | Expr.app .. => do 68 | let e' ← etaExpand e 69 | if e'.isApp then 70 | e'.withApp fun f args => do 71 | return mkAppN f (← args.mapM etaLong) 72 | else 73 | etaLong e 74 | | Expr.proj _ _ b => return e.updateProj! (← etaLong b) 75 | | Expr.mdata _ b => etaLong b 76 | | _ => etaExpand e 77 | where 78 | mkForallFVarsEtaLong (xs : Array Expr) (b : Expr) : MetaM Expr := do 79 | let mut ret := b 80 | for x in xs.reverse do 81 | ret ← mkForallFVars #[x] ret 82 | -- Eta-expand type 83 | let ty ← etaLong (ret.bindingDomain!) 84 | ret := ret.updateForall! ret.bindingInfo! ty ret.bindingBody! 85 | return ret 86 | mkLambdaFVarsEtaLong (xs : Array Expr) (b : Expr) : MetaM Expr := do 87 | let mut ret := b 88 | for x in xs.reverse do 89 | ret ← mkLambdaFVars #[x] ret 90 | -- Eta-expand type 91 | let ty ← etaLong (ret.bindingDomain!) 92 | ret := ret.updateLambda! ret.bindingInfo! ty ret.bindingBody! 93 | return ret 94 | 95 | /-- Applies eta reduction to `e` and all of its subexpressions 96 | TODO: Make sure that when calling this function from other files, 97 | the term supplied to it does not contain loose bound variables. -/ 98 | partial def etaReduce (e : Expr) : MetaM Expr := do 99 | let post (e : Expr) : MetaM TransformStep := do 100 | match e.etaExpanded? with 101 | | some f => return .done f 102 | | none => return .done e 103 | Meta.transform e (post := post) (usedLetOnly := true) 104 | 105 | /-- Instantiates mvars then applies beta and eta reduction exhaustively. -/ 106 | def betaEtaReduceInstMVars (e : Expr) : MetaM Expr := do 107 | let e ← instantiateMVars e 108 | let e ← Core.betaReduce e 109 | etaReduce e 110 | 111 | /-- Applies beta and eta reduction exhaustively -/ 112 | def betaEtaReduce (e : Expr) : MetaM Expr := do 113 | let e ← Core.betaReduce e 114 | etaReduce e 115 | -------------------------------------------------------------------------------- /Duper/Util/RestoreCoreMCodeGen.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | import Duper.TPTP 3 | import Duper.Util.InspectCoreM 4 | import Lean 5 | 6 | open Lean 7 | 8 | def codegen : CoreM Unit := do 9 | let mut file := #[] 10 | file := file.push "import Lean" 11 | file := file.push "import Duper.Tactic" 12 | file := file.push "import Duper.TPTP" 13 | file := file.push "open Lean" 14 | file := file.push "-- Desperate time, desperate measures" 15 | file := file.push "-- This file is generated by RestoreCoreMCodeGen.lean" 16 | file := file.push "" 17 | file := file.push "-- Names being added:" 18 | let names ← Duper.inspectCoreM "Duper" 19 | for name in names do 20 | file := file.push (" -- " ++ name.toString) 21 | file := file.push (← Duper.addNamesToCoreM "restoreCoreM" names) 22 | IO.FS.writeFile "Duper/Util/RestoreCoreM.lean" (String.intercalate "\n" file.toList) 23 | 24 | #eval codegen -------------------------------------------------------------------------------- /DuperOnMathlib/.gitignore: -------------------------------------------------------------------------------- 1 | /build 2 | /lake-packages/* 3 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib.lean: -------------------------------------------------------------------------------- 1 | import DuperOnMathlib.Prime 2 | import DuperOnMathlib.Cantor 3 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/Cantor.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Nat.Basic 2 | import Mathlib.Tactic 3 | import DuperOnMathlib.DuperInterface 4 | 5 | open Function 6 | 7 | variable {α : Type*} 8 | 9 | theorem Cantor : ∀ f : α → Set α, ¬ Surjective f := by 10 | intro f Surjf 11 | rcases Surjf {i | i ∉ f i} with ⟨a, h⟩ 12 | have : ¬ a ∉ f a := by 13 | -- nice 14 | duper [h, Set.mem_setOf_eq] 15 | apply this 16 | rw [h] 17 | exact this 18 | 19 | theorem Cantor2 : ∀ f : α → Set α, ¬ Surjective f := by 20 | intro f Surjf 21 | rcases Surjf {i | i ∉ f i} with ⟨a, h⟩ 22 | have : a ∈ f a ↔ a ∉ f a := by 23 | duper [h, Set.mem_setOf_eq] 24 | duper [this] 25 | 26 | -- Nice! 27 | theorem Cantor3 : ∀ f : α → Set α, ¬ Surjective f := by 28 | intro f Surjf 29 | have := Surjf {i | i ∉ f i} 30 | duper [Set.mem_setOf_eq, this] 31 | 32 | namespace original 33 | 34 | theorem Cantor : ∀ f : α → Set α, ¬ Surjective f := by 35 | intro f Surjf 36 | rcases Surjf {i | i ∉ f i} with ⟨a, h⟩ 37 | have : ¬ a ∉ f a := by 38 | intro h' 39 | apply h' 40 | rw [h] 41 | exact h' 42 | apply this 43 | rw [h] 44 | exact this 45 | 46 | theorem Cantor_alt : ∀ f : α → Set α, ¬ Surjective f := by 47 | intro f Surjf 48 | rcases Surjf {i | i ∉ f i} with ⟨a, h⟩ 49 | have : a ∈ f a ↔ a ∉ f a := by 50 | nth_rewrite 1 [h]; rfl 51 | tauto 52 | 53 | end original 54 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/DuperInterface.lean: -------------------------------------------------------------------------------- 1 | import Auto.Tactic 2 | import Duper.Tactic 3 | 4 | open Lean Auto 5 | 6 | def Auto.duperRaw (lemmas : Array Lemma) : MetaM Expr := do 7 | /- Adding `isFromGoal := false` to each formula because there is no means of distinguishing goal formulas 8 | from non-goal formulas in this context -/ 9 | let lemmas : Array (Expr × Expr × Array Name × Bool) ← 10 | lemmas.mapM (fun ⟨proof, ty, _⟩ => do return (ty, ← Meta.mkAppM ``eq_true #[proof], #[], false)) 11 | runDuper lemmas.data 0 12 | 13 | def Auto.duperPort (lemmas : Array Lemma) : MetaM Expr := do 14 | /- Adding `isFromGoal := false` to each formula because there is no means of distinguishing goal formulas 15 | from non-goal formulas in this context -/ 16 | let formulas : Array (Expr × Expr × Array Name × Bool) ← 17 | lemmas.mapM (fun ⟨proof, ty, _⟩ => do return (ty, ← Meta.mkAppM ``eq_true #[proof], #[], false)) 18 | runDuperPortfolioMode formulas.data .none 19 | { portfolioMode := true, 20 | portfolioInstance := none, 21 | inhabitationReasoning := none, 22 | preprocessing := PreprocessingOption.NoPreprocessing, -- It would be redundant to enable Auto's preprocessing for Auto calls 23 | includeExpensiveRules := none, 24 | selFunction := none 25 | } 26 | .none 27 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/HigherOrder.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | import Mathlib.Algebra.BigOperators.Ring 3 | import Duper.Tactic 4 | 5 | open BigOperators 6 | open Finset 7 | 8 | example (s : Finset ℝ) : ∑ i in s, (i + 1) = ∑ i in s, i + ∑ i in s, 1 := by 9 | rw [sum_add_distrib] 10 | 11 | example (s : Finset ℝ) : ∑ i in s, (i + 1) = ∑ i in s, i + ∑ i in s, 1 := by 12 | duper [sum_add_distrib] 13 | 14 | example (s : Finset ℝ) : ∑ i in s, (i + 1) = ∑ i in s, i + s.card := by 15 | simp only [sum_add_distrib, sum_const, nsmul_eq_mul, mul_one] 16 | 17 | example (s : Finset ℝ) : ∑ i in s, (i + 1) = ∑ i in s, i + s.card := by 18 | duper [sum_add_distrib, sum_const, nsmul_eq_mul, mul_one] 19 | 20 | example (s : Finset ℝ) : ∑ i in s, 2 * i = 2 * ∑ i in s, i := by 21 | simp only [mul_sum] 22 | 23 | example (s : Finset ℝ) : ∑ i in s, 2 * i = 2 * ∑ i in s, i := by 24 | duper [mul_sum] 25 | 26 | example (s : Finset ℝ) : ∑ i in s, (i^2 + 2 * i + 1) = 27 | ∑ i in s, i^2 + 2 * ∑ i in s, i + ∑ i in s, 1 := by 28 | simp only [sum_add_distrib] 29 | rw [←mul_sum] 30 | 31 | example (s : Finset ℝ) : ∑ i in s, (i^2 + 2 * i + 1) = 32 | ∑ i in s, i^2 + 2 * ∑ i in s, i + ∑ i in s, 1 := by 33 | rw [sum_add_distrib, sum_add_distrib, ←mul_sum] 34 | 35 | example (s : Finset ℝ) : ∑ i in s, (i^2 + 2 * i + 1) = 36 | ∑ i in s, i^2 + 2 * ∑ i in s, i + ∑ i in s, 1 := by 37 | duper [sum_add_distrib, @mul_sum _ _ s 2 (fun i => i)] 38 | 39 | -- fails 40 | -- example (s : Finset ℝ) : ∑ i in s, (i^2 + 2 * i + 1) = 41 | -- ∑ i in s, i^2 + 2 * ∑ i in s, i + ∑ i in s, 1 := by 42 | -- duper [sum_add_distrib, mul_sum] 43 | 44 | -- but this succeeds 45 | example (s : Finset ℝ) : ∑ i in s, (i^2 + 2 * i) = 46 | ∑ i in s, i^2 + 2 * ∑ i in s, i := by 47 | duper [sum_add_distrib, mul_sum] 48 | 49 | -- fails 50 | -- example (s : Finset ℝ) (n : ℝ) : ∑ i in s, (i^2 + 2 * i + n) = 51 | -- ∑ i in s, i^2 + 2 * ∑ i in s, i + ∑ i in s, n := by 52 | -- duper [sum_add_distrib, mul_sum] 53 | 54 | -- fails 55 | -- example (s : Finset ℝ) (n : ℝ) : ∑ i in s, (n + i^2 + 2 * i) = 56 | -- ∑ i in s, n + ∑ i in s, i^2 + 2 * ∑ i in s, i := by 57 | -- duper [sum_add_distrib, mul_sum] 58 | 59 | -- fails 60 | -- example (s : Finset ℝ) (n : ℝ) : ∑ i in s, (i^2 + 2 * n + 1) = 61 | -- ∑ i in s, i^2 + 2 * ∑ i in s, n + ∑ i in s, 1 := by 62 | -- duper [sum_add_distrib, mul_sum] 63 | 64 | section 65 | variable {α : Type*} [CommRing α] 66 | 67 | example (s : Finset α) : ∑ i in s, (i + 1) = ∑ i in s, i + ∑ i in s, 1 := by 68 | rw [sum_add_distrib] 69 | 70 | example (s : Finset α) : ∑ i in s, (i + 1) = ∑ i in s, i + ∑ i in s, 1 := by 71 | duper [sum_add_distrib] 72 | 73 | example (s : Finset α) : ∑ i in s, (i + 1) = ∑ i in s, i + s.card := by 74 | simp only [sum_add_distrib, sum_const, nsmul_eq_mul, mul_one] 75 | 76 | example (s : Finset α) : ∑ i in s, (i + 1) = ∑ i in s, i + s.card := by 77 | duper [sum_add_distrib, sum_const, nsmul_eq_mul, mul_one] 78 | 79 | example (s : Finset α) : ∑ i in s, 2 * i = 2 * ∑ i in s, i := by 80 | simp only [mul_sum] 81 | 82 | example (s : Finset α) : ∑ i in s, 2 * i = 2 * ∑ i in s, i := by 83 | duper [mul_sum] 84 | 85 | example (s : Finset α) : ∑ i in s, (i^2 + 2 * i + 1) = 86 | ∑ i in s, i^2 + 2 * ∑ i in s, i + ∑ i in s, 1 := by 87 | simp only [sum_add_distrib] 88 | rw [←mul_sum] 89 | 90 | example (s : Finset α) : ∑ i in s, (i^2 + 2 * i + 1) = 91 | ∑ i in s, i^2 + 2 * ∑ i in s, i + ∑ i in s, 1 := by 92 | rw [sum_add_distrib, sum_add_distrib, ←mul_sum] 93 | 94 | example (s : Finset α) : ∑ i in s, (i^2 + 2 * i + 1) = 95 | ∑ i in s, i^2 + 2 * ∑ i in s, i + ∑ i in s, 1 := by 96 | duper [sum_add_distrib, @mul_sum _ _ s 2 (fun i => i)] 97 | 98 | --fails 99 | -- example (s : Finset α) : ∑ i in s, (i^2 + 2 * i + 1) = 100 | -- ∑ i in s, i^2 + 2 * ∑ i in s, i + ∑ i in s, 1 := by 101 | -- duper [sum_add_distrib, mul_sum] 102 | 103 | 104 | end 105 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/IVT.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | 3 | def ApproachesAt (f : ℝ → ℝ) (b : ℝ) (a : ℝ) := 4 | ∀ ε, 0 < ε → ∃ δ, 0 < δ ∧ ∀ x, abs (x - a) < δ → abs (f x - b) < ε 5 | 6 | def Continuous (f : ℝ → ℝ) := ∀ x, ApproachesAt f (f x) x 7 | 8 | theorem continuous_id : Continuous (λ x ↦ x) := by 9 | intro x 10 | intro ε εpos 11 | use ε 12 | constructor 13 | exact εpos 14 | intro x' h 15 | simp [h] 16 | 17 | theorem ivt {f : ℝ → ℝ} {a b : ℝ} (aleb : a ≤ b) 18 | (ctsf : Continuous f) (hfa : f a < 0) (hfb : 0 < f b) : 19 | ∃ x, a ≤ x ∧ x ≤ b ∧ f x = 0 := by 20 | set S := { x | f x ≤ 0 ∧ a ≤ x ∧ x ≤ b } with hS 21 | have aS : a ∈ S := by 22 | exact ⟨le_of_lt hfa, le_refl a, aleb⟩ 23 | have nonemptyS : S.Nonempty := ⟨a, aS⟩ 24 | have bddS : BddAbove S := by 25 | use b 26 | intro x xS 27 | exact xS.2.2 28 | set x₀ := sSup S with hx₀ 29 | have alex₀ : a ≤ x₀ := le_csSup bddS aS 30 | have x₀leb : x₀ ≤ b := by 31 | apply csSup_le nonemptyS 32 | intro x xS 33 | exact xS.2.2 34 | have : ¬ 0 < f x₀ := by 35 | intro h 36 | rcases ctsf x₀ _ h with ⟨δ, δpos, hδ⟩ 37 | set x' := x₀ - δ / 2 with hx' 38 | have : ∃ x'' ∈ S, x' < x'' := by 39 | apply exists_lt_of_lt_csSup nonemptyS 40 | rw [hx']; linarith 41 | rcases this with ⟨x'', x''S, x'lt⟩ 42 | have : x'' ≤ x₀ := by 43 | exact le_csSup bddS x''S 44 | have : abs (f x'' - f x₀) < f x₀ := by 45 | apply hδ 46 | rw [abs_of_nonpos, neg_sub] 47 | simp [hx₀] at * 48 | linarith 49 | linarith 50 | rw [abs_lt] at this 51 | simp [hS] at x''S 52 | linarith 53 | have : ¬ f x₀ < 0 := by 54 | intro h 55 | have : 0 < - f x₀ := by 56 | linarith 57 | rcases ctsf x₀ _ this with ⟨δ, δpos, hδ⟩ 58 | have x₀lt1 : x₀ < b := by 59 | apply lt_of_le_of_ne x₀leb 60 | intro h' 61 | rw [h'] at h 62 | exact lt_asymm hfb h 63 | have minpos : 0 < min (δ / 2) (b - x₀) := by 64 | apply lt_min; linarith 65 | linarith 66 | set x' := x₀ + min (δ / 2) (b - x₀) with hx' 67 | have : abs (f x' - f x₀) < - f x₀ := by 68 | apply hδ 69 | simp [hx'] 70 | rw [abs_of_pos minpos] 71 | apply lt_of_le_of_lt (min_le_left _ _) 72 | linarith 73 | have x'S : x' ∈ S := by 74 | dsimp [hS, hx'] 75 | rw [abs_lt] at this 76 | constructor; linarith 77 | constructor; linarith 78 | linarith [min_le_right (δ / 2) (b - x₀)] 79 | have : x' ≤ x₀ := le_csSup bddS x'S 80 | dsimp [hx'] at this 81 | linarith 82 | have : f x₀ = 0 := by linarith 83 | use x₀, alex₀, x₀leb, this 84 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/ListExamples.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.List.Basic 2 | 3 | section 4 | variable {α : Type*} [LinearOrder α] 5 | 6 | /-- Returns a pair consisting of a minimal element of the list (or ⊥ if the list is empty) and the 7 | maximal element element of the list (or ⊤ if the list if empty) -/ 8 | def foo : List α → WithBot α × WithTop α 9 | | [] => (⊥, ⊤) 10 | | (x :: xs) => 11 | let (min, max) := foo xs 12 | (min ⊓ x, max ⊔ x) 13 | 14 | /-- An alternative definition, using `foldr` -/ 15 | def foo' (xs : List α) : WithBot α × WithTop α := 16 | xs.foldr (fun (x : α) (p : WithBot α × WithTop α) => (Prod.fst p ⊓ x, Prod.snd p ⊔ x)) (⊥, ⊤) 17 | 18 | theorem foo'_cons (x : α) (xs : List α) : foo' (x :: xs) = ((foo' xs).fst ⊓ x, (foo' xs).snd ⊔ x) := by 19 | simp [foo'] 20 | 21 | theorem foo_eq_foo' (xs : List α) : foo xs = foo' xs := by 22 | induction xs with 23 | | nil => rfl 24 | | cons x xs ih => 25 | simp [foo, foo'_cons, ih] 26 | 27 | theorem fst_foo_le (xs : List α) : ∀ x ∈ xs, (foo xs).fst ≤ x := by 28 | induction xs with 29 | | nil => simp [foo] 30 | | cons x xs ih => 31 | intro y h 32 | simp [foo] 33 | simp at h 34 | cases h with 35 | | inl h => 36 | simp [h] 37 | | inr h => 38 | simp [ih _ h] 39 | 40 | end 41 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/MIL/C02_Basics/S01_Calculating.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | import Duper.Tactic 3 | 4 | -- An example. 5 | example (a b c : ℝ) : a * b * c = b * (a * c) := by 6 | duper [mul_comm, mul_assoc] 7 | 8 | -- Try these. 9 | example (a b c : ℝ) : c * b * a = b * (a * c) := by 10 | duper [mul_comm, mul_assoc] 11 | 12 | example (a b c : ℝ) : a * (b * c) = b * (a * c) := by 13 | duper [mul_comm, mul_assoc] 14 | 15 | -- An example. 16 | example (a b c : ℝ) : a * b * c = b * c * a := by 17 | duper [mul_comm, mul_assoc] 18 | /- Try doing the first of these without providing any arguments at all, 19 | and the second with only one argument. -/ 20 | example (a b c : ℝ) : a * (b * c) = b * (c * a) := by 21 | duper [mul_comm, mul_assoc] 22 | 23 | example (a b c : ℝ) : a * (b * c) = b * (a * c) := by 24 | duper [mul_comm, mul_assoc] 25 | 26 | -- Using facts from the local context. 27 | example (a b c d e f : ℝ) (h : a * b = c * d) (h' : e = f) : a * (b * e) = c * (d * f) := by 28 | duper [mul_comm, mul_assoc, h, h'] 29 | 30 | example (a b c d e f : ℝ) (h : b * c = e * f) : a * b * c * d = a * e * f * d := by 31 | duper [mul_comm, mul_assoc, h] 32 | 33 | example (a b c d : ℝ) (hyp : c = b * a - d) (hyp' : d = a * b) : c = 0 := by 34 | duper [mul_comm, mul_assoc, hyp, hyp', sub_self] 35 | 36 | example (a b c d e f : ℝ) (h : a * b = c * d) (h' : e = f) : a * (b * e) = c * (d * f) := by 37 | duper [mul_assoc, h, h'] 38 | 39 | section 40 | 41 | variable (a b c d e f : ℝ) 42 | 43 | example (h : a * b = c * d) (h' : e = f) : a * (b * e) = c * (d * f) := by 44 | duper [mul_assoc, h, h'] 45 | 46 | end 47 | 48 | section 49 | variable (a b c : ℝ) 50 | 51 | #check a 52 | #check a + b 53 | #check (a : ℝ) 54 | #check mul_comm a b 55 | #check (mul_comm a b : a * b = b * a) 56 | #check mul_assoc c a b 57 | #check mul_comm a 58 | #check mul_comm 59 | 60 | end 61 | 62 | section 63 | variable (a b : ℝ) 64 | 65 | example : (a + b) * (a + b) = a * a + 2 * (a * b) + b * b := by 66 | -- duper needs help here 67 | simp only [mul_add, add_mul] 68 | -- if we add `add_mul` to the next line, duper times out 69 | duper [add_assoc, mul_comm, two_mul] 70 | 71 | end 72 | 73 | -- Try these. For the second, use the theorems listed underneath. 74 | section 75 | variable (a b c d : ℝ) 76 | 77 | example : (a + b) * (c + d) = a * c + a * d + b * c + b * d := by 78 | duper [add_mul, mul_add, add_comm, mul_comm, add_assoc] 79 | 80 | example (a b : ℝ) : (a + b) * (a - b) = a ^ 2 - b ^ 2 := by 81 | -- simp [add_mul, mul_add, mul_sub, pow_two, mul_comm, ←sub_sub] 82 | -- simp only [add_mul, mul_add, mul_sub, sub_mul, pow_two, ←sub_sub, mul_comm, add_sub_cancel] 83 | -- duper [add_mul, mul_add, mul_sub, sub_mul, pow_two, sub_sub, mul_comm, add_sub_cancel] 84 | simp only [add_mul, mul_add, mul_sub, sub_mul] 85 | -- trying to move any of the above to duper fails 86 | duper [pow_two, sub_sub, add_sub_cancel, mul_comm] 87 | 88 | #check pow_two a 89 | #check mul_sub a b c 90 | #check add_mul a b c 91 | #check add_sub a b c 92 | #check sub_sub a b c 93 | #check add_zero a 94 | 95 | end 96 | 97 | -- Examples. 98 | 99 | section 100 | variable (a b c d : ℝ) 101 | 102 | example (a b c d : ℝ) (hyp : c = d * a + b) (hyp' : b = a * d) : c = 2 * a * d := by 103 | duper [hyp, hyp', mul_comm d a, two_mul, mul_assoc] 104 | /- 105 | rw [hyp'] at hyp 106 | rw [mul_comm d a] at hyp 107 | rw [← two_mul (a * d)] at hyp 108 | rw [← mul_assoc 2 a d] at hyp 109 | exact hyp 110 | -/ 111 | 112 | example : c * b * a = b * (a * c) := by 113 | ring 114 | 115 | example : (a + b) * (a + b) = a * a + 2 * (a * b) + b * b := by 116 | ring 117 | 118 | example : (a + b) * (a - b) = a ^ 2 - b ^ 2 := by 119 | ring 120 | 121 | example (hyp : c = d * a + b) (hyp' : b = a * d) : c = 2 * a * d := by 122 | rw [hyp, hyp'] 123 | ring 124 | 125 | end 126 | 127 | example (a b c : ℕ) (h : a + b = c) : (a + b) * (a + b) = a * c + b * c := by 128 | duper [h, add_mul] 129 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/MIL/C02_Basics/S05_Proving_Facts_about_Algebraic_Structures.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Topology.MetricSpace.Basic 2 | import Duper.Tactic 3 | 4 | section 5 | variable {α : Type*} [Lattice α] 6 | variable (x y z : α) 7 | 8 | #check x ⊓ y 9 | #check (inf_le_left : x ⊓ y ≤ x) 10 | #check (inf_le_right : x ⊓ y ≤ y) 11 | #check (le_inf : z ≤ x → z ≤ y → z ≤ x ⊓ y) 12 | #check x ⊔ y 13 | #check (le_sup_left : x ≤ x ⊔ y) 14 | #check (le_sup_right : y ≤ x ⊔ y) 15 | #check (sup_le : x ≤ z → y ≤ z → x ⊔ y ≤ z) 16 | 17 | example : x ⊓ y = y ⊓ x := by 18 | duper [le_antisymm_iff, le_inf, inf_le_left, inf_le_right] 19 | 20 | example : x ⊓ y ⊓ z = x ⊓ (y ⊓ z) := by 21 | apply le_antisymm 22 | . apply le_inf 23 | . duper [le_antisymm_iff, le_inf, inf_le_left, inf_le_right, le_trans] 24 | . duper [le_antisymm_iff, le_inf, inf_le_left, inf_le_right, le_trans] 25 | . apply le_inf 26 | . duper [le_antisymm_iff, le_inf, inf_le_left, inf_le_right, le_trans] 27 | . duper [le_antisymm_iff, le_inf, inf_le_left, inf_le_right, le_trans] 28 | 29 | example : x ⊔ y = y ⊔ x := by 30 | duper [le_antisymm_iff, sup_le, le_sup_left, le_sup_right] 31 | 32 | example : x ⊔ y ⊔ z = x ⊔ (y ⊔ z) := by 33 | apply le_antisymm 34 | . apply sup_le 35 | . duper [le_antisymm_iff, sup_le, le_sup_left, le_sup_right, le_trans] 36 | . duper [le_antisymm_iff, sup_le, le_sup_left, le_sup_right, le_trans] 37 | . apply sup_le 38 | . duper [le_antisymm_iff, sup_le, le_sup_left, le_sup_right, le_trans] 39 | . duper [le_antisymm_iff, sup_le, le_sup_left, le_sup_right, le_trans] 40 | 41 | theorem absorb1 : x ⊓ (x ⊔ y) = x := by 42 | apply le_antisymm 43 | · apply inf_le_left 44 | apply le_inf 45 | · apply le_refl 46 | apply le_sup_left 47 | 48 | theorem absorb1' : x ⊓ (x ⊔ y) = x := by 49 | duper [le_antisymm, inf_le_left, le_inf, le_refl, le_sup_left] 50 | 51 | theorem absorb2 : x ⊔ x ⊓ y = x := by 52 | apply le_antisymm 53 | · apply sup_le 54 | · apply le_refl 55 | apply inf_le_left 56 | apply le_sup_left 57 | 58 | theorem absorb2' : x ⊔ x ⊓ y = x := by 59 | duper [le_antisymm, le_sup_left, sup_le, le_refl, inf_le_left] 60 | 61 | end 62 | 63 | section 64 | variable {α : Type*} [DistribLattice α] 65 | variable (x y z : α) 66 | 67 | #check (inf_sup_left : x ⊓ (y ⊔ z) = x ⊓ y ⊔ x ⊓ z) 68 | #check (inf_sup_right : (x ⊔ y) ⊓ z = x ⊓ z ⊔ y ⊓ z) 69 | #check (sup_inf_left : x ⊔ y ⊓ z = (x ⊔ y) ⊓ (x ⊔ z)) 70 | #check (sup_inf_right : x ⊓ y ⊔ z = (x ⊔ z) ⊓ (y ⊔ z)) 71 | end 72 | 73 | section 74 | variable {α : Type*} [Lattice α] 75 | variable (a b c : α) 76 | 77 | example (h : ∀ x y z : α, x ⊓ (y ⊔ z) = x ⊓ y ⊔ x ⊓ z) : a ⊔ b ⊓ c = (a ⊔ b) ⊓ (a ⊔ c) := by 78 | rw [h, @inf_comm _ _ (a ⊔ b), absorb1, @inf_comm _ _ (a ⊔ b), h, ← sup_assoc, @inf_comm _ _ c a, 79 | absorb2, inf_comm] 80 | 81 | example (h : ∀ x y z : α, x ⊓ (y ⊔ z) = x ⊓ y ⊔ x ⊓ z) : a ⊔ b ⊓ c = (a ⊔ b) ⊓ (a ⊔ c) := by 82 | rw [h, @inf_comm _ _ (a ⊔ b), absorb1, @inf_comm _ _ (a ⊔ b), h, ← sup_assoc, @inf_comm _ _ c a] 83 | -- If we add the redundant lemma `sup_assoc` to this, it fails. 84 | duper [h, absorb1, absorb2, inf_comm] 85 | 86 | example (h : ∀ x y z : α, x ⊔ y ⊓ z = (x ⊔ y) ⊓ (x ⊔ z)) : a ⊓ (b ⊔ c) = a ⊓ b ⊔ a ⊓ c := by 87 | rw [h, @sup_comm _ _ (a ⊓ b), absorb2, @sup_comm _ _ (a ⊓ b), h, ← inf_assoc, @sup_comm _ _ c a, 88 | absorb1, sup_comm] 89 | 90 | end 91 | 92 | section 93 | variable {R : Type*} [StrictOrderedRing R] 94 | variable (a b c : R) 95 | 96 | #check (add_le_add_left : a ≤ b → ∀ c, c + a ≤ c + b) 97 | #check (mul_pos : 0 < a → 0 < b → 0 < a * b) 98 | 99 | #check (mul_nonneg : 0 ≤ a → 0 ≤ b → 0 ≤ a * b) 100 | 101 | -- First, the manual version 102 | 103 | theorem aux1 (h : a ≤ b) : 0 ≤ b - a := by 104 | rw [← sub_self a, sub_eq_add_neg, sub_eq_add_neg, add_comm, add_comm b] 105 | apply add_le_add_left h 106 | 107 | theorem aux2 (h : 0 ≤ b - a) : a ≤ b := by 108 | rw [← add_zero a, ← sub_add_cancel b a, add_comm (b - a)] 109 | apply add_le_add_left h 110 | 111 | example (h : a ≤ b) (h' : 0 ≤ c) : a * c ≤ b * c := by 112 | have h1 : 0 ≤ (b - a) * c := mul_nonneg (aux1 _ _ h) h' 113 | rw [sub_mul] at h1 114 | exact aux2 _ _ h1 115 | 116 | -- Now, with duper 117 | 118 | example (h : a ≤ b) : 0 ≤ b - a := by 119 | duper [sub_self, sub_eq_add_neg, add_comm, add_comm, add_le_add_left, h] 120 | 121 | example (h : 0 ≤ b - a) : a ≤ b := by 122 | duper [add_zero, sub_add_cancel, add_comm, add_le_add_left, h] 123 | 124 | example (h : a ≤ b) (h' : 0 ≤ c) : a * c ≤ b * c := by 125 | duper [mul_nonneg, aux1, aux2, sub_mul, *] 126 | 127 | -- Ambitious: 128 | 129 | example (h : a ≤ b) (h' : 0 ≤ c) : a * c ≤ b * c := by 130 | -- try replacing `aux2` with the facts used to prove it 131 | -- duper [mul_nonneg, aux1, add_zero, sub_add_cancel, add_comm, add_le_add_left, sub_mul, *] 132 | -- try replacing `aux1` with the facts used to prove it 133 | -- duper [mul_nonneg, sub_self, sub_eq_add_neg, add_comm, add_comm, add_le_add_left, aux2, sub_mul, *] 134 | sorry 135 | 136 | end 137 | 138 | section 139 | variable {X : Type*} [MetricSpace X] 140 | variable (x y z : X) 141 | 142 | #check (dist_self x : dist x x = 0) 143 | #check (dist_comm x y : dist x y = dist y x) 144 | #check (dist_triangle x y z : dist x z ≤ dist x y + dist y z) 145 | 146 | -- manual proof 147 | example (x y : X) : 0 ≤ dist x y :=by 148 | have : 0 ≤ dist x y + dist y x := by 149 | rw [← dist_self x] 150 | apply dist_triangle 151 | linarith [dist_comm x y] 152 | 153 | example (x y : X) : 0 ≤ dist x y := by 154 | have : 0 ≤ dist x y + dist y x := by 155 | duper [dist_self, dist_triangle] 156 | linarith [dist_comm x y] 157 | 158 | end 159 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/MIL/C03_Logic/S01_Implication_and_the_Universal_Quantifier.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | import Duper.Tactic 3 | 4 | namespace C03S01 5 | 6 | def FnUb (f : ℝ → ℝ) (a : ℝ) : Prop := 7 | ∀ x, f x ≤ a 8 | 9 | def FnLb (f : ℝ → ℝ) (a : ℝ) : Prop := 10 | ∀ x, a ≤ f x 11 | 12 | section 13 | variable (f g : ℝ → ℝ) (a b : ℝ) 14 | 15 | example (hfa : FnUb f a) (hgb : FnUb g b) : FnUb (fun x ↦ f x + g x) (a + b) := by 16 | duper [add_le_add, *, FnUb] 17 | 18 | example (hfa : FnLb f a) (hgb : FnLb g b) : FnLb (fun x ↦ f x + g x) (a + b) := by 19 | duper [add_le_add, *, FnLb] 20 | 21 | example (nnf : FnLb f 0) (nng : FnLb g 0) : FnLb (fun x ↦ f x * g x) 0 := by 22 | duper [mul_nonneg, *, FnLb] 23 | 24 | example (hfa : FnUb f a) (hgb : FnUb g b) (nng : FnLb g 0) (nna : 0 ≤ a) : 25 | FnUb (fun x ↦ f x * g x) (a * b) := by 26 | duper [mul_le_mul, *, FnUb, FnLb] 27 | 28 | end 29 | 30 | section 31 | variable {α : Type*} {R : Type*} [OrderedCancelAddCommMonoid R] 32 | 33 | #check add_le_add 34 | 35 | def FnUb' (f : α → R) (a : R) : Prop := 36 | ∀ x, f x ≤ a 37 | 38 | theorem fnUb_add {f g : α → R} {a b : R} (hfa : FnUb' f a) (hgb : FnUb' g b) : 39 | FnUb' (fun x ↦ f x + g x) (a + b) := fun x ↦ add_le_add (hfa x) (hgb x) 40 | 41 | end 42 | 43 | example (f : ℝ → ℝ) (h : Monotone f) : ∀ {a b}, a ≤ b → f a ≤ f b := 44 | @h 45 | 46 | section 47 | variable (f g : ℝ → ℝ) 48 | 49 | example (mf : Monotone f) (mg : Monotone g) : Monotone fun x ↦ f x + g x := by 50 | -- interesting: duper doesn't work unless you intro first 51 | --intro a b aleb 52 | rw [Monotone] 53 | duper [Monotone, add_le_add, *] 54 | 55 | example {c : ℝ} (mf : Monotone f) (nnc : 0 ≤ c) : Monotone fun x ↦ c * f x := by 56 | intro a b aleb 57 | apply mul_le_mul_of_nonneg_left _ nnc 58 | apply mf aleb 59 | 60 | example {c : ℝ} (mf : Monotone f) (nnc : 0 ≤ c) : Monotone fun x ↦ c * f x := by 61 | duper [*, Monotone, mul_le_mul_of_nonneg_left] 62 | 63 | example (mf : Monotone f) (mg : Monotone g) : Monotone fun x ↦ f (g x) := by 64 | intro a b aleb 65 | apply mf 66 | apply mg 67 | apply aleb 68 | 69 | example (mf : Monotone f) (mg : Monotone g) : Monotone fun x ↦ f (g x) := by 70 | duper [Monotone, *] 71 | 72 | def FnEven (f : ℝ → ℝ) : Prop := 73 | ∀ x, f x = f (-x) 74 | 75 | def FnOdd (f : ℝ → ℝ) : Prop := 76 | ∀ x, f x = -f (-x) 77 | 78 | example (ef : FnEven f) (eg : FnEven g) : FnEven fun x ↦ f x + g x := by 79 | intro x 80 | calc 81 | (fun x ↦ f x + g x) x = f x + g x := rfl 82 | _ = f (-x) + g (-x) := by rw [ef, eg] 83 | 84 | example (ef : FnEven f) (eg : FnEven g) : FnEven fun x ↦ f x + g x := by 85 | duper [*, FnEven] 86 | 87 | example (of : FnOdd f) (og : FnOdd g) : FnEven fun x ↦ f x * g x := by 88 | duper [*, FnOdd, FnEven, neg_mul_neg] 89 | 90 | example (ef : FnEven f) (og : FnOdd g) : FnOdd fun x ↦ f x * g x := by 91 | intro x 92 | dsimp 93 | rw [ef, og, neg_mul_eq_mul_neg] 94 | 95 | example (ef : FnEven f) (og : FnOdd g) : FnEven fun x ↦ f (g x) := by 96 | intro x 97 | dsimp 98 | rw [og, ← ef] 99 | 100 | example (ef : FnEven f) (og : FnOdd g) : FnOdd fun x ↦ f x * g x := by 101 | duper [*, FnEven, FnOdd, neg_mul_eq_mul_neg] 102 | 103 | example (ef : FnEven f) (og : FnOdd g) : FnEven fun x ↦ f (g x) := by 104 | duper [*, FnEven, FnOdd] 105 | 106 | end 107 | 108 | section 109 | 110 | variable {α : Type*} (r s t : Set α) 111 | 112 | example : s ⊆ s := by 113 | intro x xs 114 | exact xs 115 | 116 | example : s ⊆ s := by 117 | duper [Set.subset_def] 118 | 119 | theorem Subset.refl : s ⊆ s := fun x xs ↦ xs 120 | 121 | theorem Subset.trans : r ⊆ s → s ⊆ t → r ⊆ t := by 122 | duper [Set.subset_def] 123 | 124 | end 125 | 126 | section 127 | variable {α : Type*} [PartialOrder α] 128 | variable (s : Set α) (a b : α) 129 | 130 | def SetUb (s : Set α) (a : α) := 131 | ∀ x, x ∈ s → x ≤ a 132 | 133 | example (h : SetUb s a) (h' : a ≤ b) : SetUb s b := by 134 | duper [*, SetUb, le_trans] 135 | 136 | end 137 | 138 | section 139 | 140 | open Function 141 | 142 | example (c : ℝ) : Injective fun x ↦ x + c := by 143 | intro x₁ x₂ h' 144 | exact (add_left_inj c).mp h' 145 | 146 | example (c : ℝ) : Injective fun x ↦ x + c := by 147 | duper [*, Injective, add_left_inj] 148 | 149 | example {c : ℝ} (h : c ≠ 0) : Injective fun x ↦ c * x := by 150 | duper [*, Injective, mul_right_inj'] 151 | 152 | variable {α : Type*} {β : Type*} {γ : Type*} 153 | variable {g : β → γ} {f : α → β} 154 | 155 | example (injg : Injective g) (injf : Injective f) : Injective fun x ↦ g (f x) := by 156 | duper [*, Injective] 157 | 158 | end 159 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/MIL/C03_Logic/S02_The_Existential_Quantifier.lean: -------------------------------------------------------------------------------- 1 | 2 | import Mathlib.Data.Real.Basic 3 | import Duper.Tactic 4 | 5 | set_option autoImplicit true 6 | 7 | namespace C03S02 8 | 9 | example : ∃ x : ℝ, 2 < x ∧ x < 3 := by 10 | use 5 / 2 11 | norm_num 12 | 13 | 14 | def FnUb (f : ℝ → ℝ) (a : ℝ) : Prop := 15 | ∀ x, f x ≤ a 16 | 17 | def FnLb (f : ℝ → ℝ) (a : ℝ) : Prop := 18 | ∀ x, a ≤ f x 19 | 20 | def FnHasUb (f : ℝ → ℝ) := 21 | ∃ a, FnUb f a 22 | 23 | def FnHasLb (f : ℝ → ℝ) := 24 | ∃ a, FnLb f a 25 | 26 | theorem fnUb_add {f g : ℝ → ℝ} {a b : ℝ} (hfa : FnUb f a) (hgb : FnUb g b) : 27 | FnUb (fun x ↦ f x + g x) (a + b) := by 28 | simp only [FnUb] at * 29 | duper [FnUb, add_le_add, *] 30 | 31 | section 32 | 33 | variable {f g : ℝ → ℝ} 34 | 35 | example (ubf : FnHasUb f) (ubg : FnHasUb g) : FnHasUb fun x ↦ f x + g x := by 36 | rcases ubf with ⟨a, ubfa⟩ 37 | rcases ubg with ⟨b, ubgb⟩ 38 | use a + b 39 | apply fnUb_add ubfa ubgb 40 | 41 | example (ubf : FnHasUb f) (ubg : FnHasUb g) : FnHasUb fun x ↦ f x + g x := by 42 | duper [*, FnHasUb, FnUb, add_le_add] 43 | 44 | 45 | example (lbf : FnHasLb f) (lbg : FnHasLb g) : FnHasLb fun x ↦ f x + g x := by 46 | duper [*, FnHasLb, FnLb, add_le_add] 47 | 48 | example {c : ℝ} (ubf : FnHasUb f) (h : c ≥ 0) : FnHasUb fun x ↦ c * f x := by 49 | rcases ubf with ⟨a, lbfa⟩ 50 | use c * a 51 | intro x 52 | exact mul_le_mul_of_nonneg_left (lbfa x) h 53 | 54 | example {c : ℝ} (ubf : FnHasUb f) (h : c ≥ 0) : FnHasUb fun x ↦ c * f x := by 55 | duper [*, FnHasUb, FnUb, mul_le_mul_of_nonneg_left] 56 | 57 | example : FnHasUb f → FnHasUb g → FnHasUb fun x ↦ f x + g x := by 58 | rintro ⟨a, ubfa⟩ ⟨b, ubgb⟩ 59 | exact ⟨a + b, fnUb_add ubfa ubgb⟩ 60 | 61 | example : FnHasUb f → FnHasUb g → FnHasUb fun x ↦ f x + g x := by 62 | duper [*, FnHasUb, fnUb_add] 63 | 64 | example : FnHasUb f → FnHasUb g → FnHasUb fun x ↦ f x + g x := 65 | fun ⟨a, ubfa⟩ ⟨b, ubgb⟩ ↦ ⟨a + b, fnUb_add ubfa ubgb⟩ 66 | 67 | example : FnHasUb f → FnHasUb g → FnHasUb fun x ↦ f x + g x := by 68 | duper [*, FnHasUb, fnUb_add] 69 | 70 | end 71 | 72 | example (ubf : FnHasUb f) (ubg : FnHasUb g) : FnHasUb fun x ↦ f x + g x := by 73 | duper [*, FnHasUb, fnUb_add] 74 | 75 | 76 | section 77 | 78 | variable {α : Type*} [CommRing α] 79 | 80 | def SumOfSquares (x : α) := 81 | ∃ a b, x = a ^ 2 + b ^ 2 82 | 83 | theorem sumOfSquares_mul {x y : α} (sosx : SumOfSquares x) (sosy : SumOfSquares y) : 84 | SumOfSquares (x * y) := by 85 | rcases sosx with ⟨a, b, xeq⟩ 86 | rcases sosy with ⟨c, d, yeq⟩ 87 | rw [xeq, yeq] 88 | use a * c - b * d, a * d + b * c 89 | ring 90 | 91 | theorem sumOfSquares_mul' {x y : α} (sosx : SumOfSquares x) (sosy : SumOfSquares y) : 92 | SumOfSquares (x * y) := by 93 | rcases sosx with ⟨a, b, rfl⟩ 94 | rcases sosy with ⟨c, d, rfl⟩ 95 | use a * c - b * d, a * d + b * c 96 | ring 97 | 98 | end 99 | 100 | section 101 | variable {a b c : ℕ} 102 | 103 | example (divab : a ∣ b) (divbc : b ∣ c) : a ∣ c := by 104 | rcases divab with ⟨d, beq⟩ 105 | rcases divbc with ⟨e, ceq⟩ 106 | rw [ceq, beq] 107 | use d * e; ring 108 | 109 | example (divab : a ∣ b) (divbc : b ∣ c) : a ∣ c := by 110 | duper [dvd_def, mul_assoc, *] 111 | 112 | example (divab : a ∣ b) (divac : a ∣ c) : a ∣ b + c := by 113 | duper [dvd_def, mul_add, *] 114 | 115 | end 116 | 117 | section 118 | 119 | open Function 120 | 121 | example {c : ℝ} : Surjective fun x ↦ x + c := by 122 | intro x 123 | use x - c 124 | dsimp; ring 125 | 126 | example {c : ℝ} : Surjective fun x ↦ x + c := by 127 | duper [Surjective, sub_add_cancel] 128 | 129 | example {c : ℝ} (h : c ≠ 0) : Surjective fun x ↦ c * x := by 130 | duper [Surjective, div_mul_cancel, mul_comm, *] 131 | 132 | example (x y : ℝ) (h : x - y ≠ 0) : (x ^ 2 - y ^ 2) / (x - y) = x + y := by 133 | field_simp [h] 134 | ring 135 | 136 | example {f : ℝ → ℝ} (h : Surjective f) : ∃ x, f x ^ 2 = 4 := by 137 | rcases h 2 with ⟨x, hx⟩ 138 | use x 139 | rw [hx] 140 | norm_num 141 | 142 | example {f : ℝ → ℝ} (h : Surjective f) : ∃ x, f x ^ 2 = 4 := by 143 | have : (2 : ℝ)^2 = 4 := by norm_num 144 | duper [*, Surjective] 145 | 146 | end 147 | 148 | section 149 | open Function 150 | variable {α : Type*} {β : Type*} {γ : Type*} 151 | variable {g : β → γ} {f : α → β} 152 | 153 | example (surjg : Surjective g) (surjf : Surjective f) : Surjective fun x ↦ g (f x) := by 154 | duper [Surjective, *] 155 | 156 | end 157 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/MIL/C03_Logic/S03_Negation.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | import Duper.Tactic 3 | 4 | namespace C03S03 5 | 6 | section 7 | variable (a b : ℝ) 8 | 9 | example (h : a < b) : ¬b < a := by 10 | intro h' 11 | have : a < a := lt_trans h h' 12 | apply lt_irrefl a this 13 | 14 | example (h : a < b) : ¬b < a := by 15 | duper [lt_trans, lt_irrefl, *] 16 | 17 | def FnUb (f : ℝ → ℝ) (a : ℝ) : Prop := 18 | ∀ x, f x ≤ a 19 | 20 | def FnLb (f : ℝ → ℝ) (a : ℝ) : Prop := 21 | ∀ x, a ≤ f x 22 | 23 | def FnHasUb (f : ℝ → ℝ) := 24 | ∃ a, FnUb f a 25 | 26 | def FnHasLb (f : ℝ → ℝ) := 27 | ∃ a, FnLb f a 28 | 29 | variable (f : ℝ → ℝ) 30 | 31 | example (h : ∀ a, ∃ x, f x > a) : ¬FnHasUb f := by 32 | intro fnub 33 | rcases fnub with ⟨a, fnuba⟩ 34 | rcases h a with ⟨x, hx⟩ 35 | have : f x ≤ a := fnuba x 36 | linarith 37 | 38 | example (h : ∀ a, ∃ x, f x > a) : ¬FnHasUb f := by 39 | duper [*, FnHasUb, FnUb, gt_iff_lt, not_le] 40 | 41 | example (h : ∀ a, ∃ x, f x < a) : ¬FnHasLb f := by 42 | duper [*, FnHasLb, FnLb, gt_iff_lt, not_le] 43 | 44 | example : ¬FnHasUb fun x ↦ x := by 45 | have : ∀ x : ℝ, x < x + 1 := by simp 46 | duper [*, FnHasUb, FnUb, gt_iff_lt, not_le, this] 47 | 48 | #check (not_le_of_gt : a > b → ¬a ≤ b) 49 | #check (not_lt_of_ge : a ≥ b → ¬a < b) 50 | #check (lt_of_not_ge : ¬a ≥ b → a < b) 51 | #check (le_of_not_gt : ¬a > b → a ≤ b) 52 | 53 | 54 | example (h : Monotone f) (h' : f a < f b) : a < b := by 55 | apply lt_of_not_ge 56 | intro h'' 57 | apply absurd h' 58 | apply not_lt_of_ge (h h'') 59 | 60 | -- Nice! 61 | example (h : Monotone f) (h' : f a < f b) : a < b := by 62 | duper [lt_of_not_ge, not_lt_of_ge, Monotone, *] 63 | 64 | example (h : a ≤ b) (h' : f b < f a) : ¬Monotone f := by 65 | intro h'' 66 | apply absurd h' 67 | apply not_lt_of_ge 68 | apply h'' h 69 | 70 | example (h : a ≤ b) (h' : f b < f a) : ¬Monotone f := by 71 | duper [Monotone, not_lt_of_ge, *] 72 | 73 | example : ¬∀ {f : ℝ → ℝ}, Monotone f → ∀ {a b}, f a ≤ f b → a ≤ b := by 74 | intro h 75 | let f := fun x : ℝ ↦ (0 : ℝ) 76 | have monof : Monotone f := by 77 | intro a b leab 78 | rfl 79 | have h' : f 1 ≤ f 0 := le_refl _ 80 | have : (1 : ℝ) ≤ 0 := h monof h' 81 | linarith 82 | 83 | example : ¬∀ {f : ℝ → ℝ}, Monotone f → ∀ {a b}, f a ≤ f b → a ≤ b := by 84 | set f := fun x : ℝ ↦ (0 : ℝ) with hf 85 | have h' : f 1 ≤ f 0 := le_refl _ 86 | have : ¬ (1 : ℝ ) ≤ 0 := by linarith 87 | duper [Monotone, *] 88 | 89 | example (x : ℝ) (h : ∀ ε > 0, x < ε) : x ≤ 0 := by 90 | apply le_of_not_gt 91 | intro h' 92 | linarith [h _ h'] 93 | 94 | example (x : ℝ) (h : ∀ ε > 0, x < ε) : x ≤ 0 := by 95 | duper [*, le_of_not_gt, lt_irrefl] 96 | 97 | end 98 | 99 | section 100 | variable {α : Type*} (P : α → Prop) (Q : Prop) 101 | 102 | example (h : ¬∃ x, P x) : ∀ x, ¬P x := by 103 | duper [h] 104 | 105 | example (h : ∀ x, ¬P x) : ¬∃ x, P x := by 106 | duper [h] 107 | 108 | example (h : ¬∀ x, P x) : ∃ x, ¬P x := by 109 | duper [h] 110 | 111 | example (h : ∃ x, ¬P x) : ¬∀ x, P x := by 112 | duper [h] 113 | 114 | example (h : ¬¬Q) : Q := by 115 | duper [h] 116 | 117 | example (h : Q) : ¬¬Q := by 118 | duper [h] 119 | 120 | end 121 | 122 | section 123 | variable (f : ℝ → ℝ) 124 | 125 | example (h : ¬FnHasUb f) : ∀ a, ∃ x, f x > a := by 126 | intro a 127 | by_contra h' 128 | apply h 129 | use a 130 | intro x 131 | apply le_of_not_gt 132 | intro h'' 133 | apply h' 134 | use x 135 | 136 | example (h : ¬FnHasUb f) : ∀ a, ∃ x, f x > a := by 137 | duper [h, le_of_not_gt, FnHasUb, FnUb] 138 | 139 | example (h : ¬∀ a, ∃ x, f x > a) : FnHasUb f := by 140 | push_neg at h 141 | exact h 142 | 143 | example (h : ¬FnHasUb f) : ∀ a, ∃ x, f x > a := by 144 | dsimp only [FnHasUb, FnUb] at h 145 | push_neg at h 146 | exact h 147 | 148 | example (h : ¬FnHasUb f) : ∀ a, ∃ x, f x > a := by 149 | duper [FnHasUb, FnUb, h, le_of_not_gt] 150 | 151 | example (h : ¬Monotone f) : ∃ x y, x ≤ y ∧ f y < f x := by 152 | duper [Monotone, h, le_of_not_gt] 153 | 154 | 155 | example (x : ℝ) (h : ∀ ε > 0, x ≤ ε) : x ≤ 0 := by 156 | contrapose! h 157 | use x / 2 158 | constructor <;> linarith 159 | 160 | end 161 | 162 | section 163 | variable (a : ℕ) 164 | 165 | example (h : 0 < 0) : a > 37 := by 166 | duper [h, lt_irrefl] 167 | 168 | end 169 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/MIL/C03_Logic/S04_Conjunction_and_Iff.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | import Mathlib.Data.Nat.Prime 3 | import Duper.Tactic 4 | 5 | namespace C03S04 6 | 7 | example {x y : ℝ} (h₀ : x ≤ y) (h₁ : ¬y ≤ x) : x ≤ y ∧ x ≠ y := by 8 | duper [*] 9 | 10 | example {x y : ℝ} (h : x ≤ y ∧ x ≠ y) : ¬y ≤ x := by 11 | duper [*, le_antisymm] 12 | 13 | example {m n : ℕ} (h : m ∣ n ∧ m ≠ n) : m ∣ n ∧ ¬n ∣ m := by 14 | rcases h with ⟨h0, h1⟩ 15 | constructor 16 | · exact h0 17 | intro h2 18 | apply h1 19 | apply Nat.dvd_antisymm h0 h2 20 | 21 | example {m n : ℕ} (h : m ∣ n ∧ m ≠ n) : m ∣ n ∧ ¬n ∣ m := by 22 | duper [*, Nat.dvd_antisymm] 23 | 24 | example : ∃ x : ℝ, 2 < x ∧ x < 4 := 25 | ⟨5 / 2, by norm_num, by norm_num⟩ 26 | 27 | example (x y : ℝ) : (∃ z : ℝ, x < z ∧ z < y) → x < y := by 28 | rintro ⟨z, xltz, zlty⟩ 29 | exact lt_trans xltz zlty 30 | 31 | example (x y : ℝ) : (∃ z : ℝ, x < z ∧ z < y) → x < y := by 32 | duper [*, lt_trans] 33 | 34 | example : ∃ m n : ℕ, 4 < m ∧ m < n ∧ n < 10 ∧ Nat.Prime m ∧ Nat.Prime n := by 35 | use 5 36 | use 7 37 | norm_num 38 | 39 | example : ∃ m n : ℕ, 4 < m ∧ m < n ∧ n < 10 ∧ Nat.Prime m ∧ Nat.Prime n := by 40 | have : 4 < 5 ∧ 5 < 7 ∧ 7 < 10 ∧ Nat.Prime 5 ∧ Nat.Prime 7 := by norm_num 41 | duper [*] 42 | 43 | example {x y : ℝ} : x ≤ y ∧ x ≠ y → x ≤ y ∧ ¬y ≤ x := by 44 | rintro ⟨h₀, h₁⟩ 45 | use h₀ 46 | exact fun h' ↦ h₁ (le_antisymm h₀ h') 47 | 48 | example {x y : ℝ} : x ≤ y ∧ x ≠ y → x ≤ y ∧ ¬y ≤ x := by 49 | duper [le_antisymm] 50 | 51 | example {x y : ℝ} (h : x ≤ y) : ¬y ≤ x ↔ x ≠ y := by 52 | constructor 53 | · contrapose! 54 | rintro rfl 55 | rfl 56 | contrapose! 57 | exact le_antisymm h 58 | 59 | example {x y : ℝ} (h : x ≤ y) : ¬y ≤ x ↔ x ≠ y := by 60 | duper [*, le_antisymm] 61 | 62 | example {x y : ℝ} : x ≤ y ∧ ¬y ≤ x ↔ x ≤ y ∧ x ≠ y := by 63 | constructor 64 | · rintro ⟨h0, h1⟩ 65 | constructor 66 | · exact h0 67 | intro h2 68 | apply h1 69 | rw [h2] 70 | rintro ⟨h0, h1⟩ 71 | constructor 72 | · exact h0 73 | intro h2 74 | apply h1 75 | apply le_antisymm h0 h2 76 | 77 | example {x y : ℝ} : x ≤ y ∧ ¬y ≤ x ↔ x ≤ y ∧ x ≠ y := by 78 | duper [le_antisymm] 79 | 80 | theorem aux {x y : ℝ} (h : x ^ 2 + y ^ 2 = 0) : x = 0 := 81 | have h' : x ^ 2 = 0 := by linarith [pow_two_nonneg x, pow_two_nonneg y] 82 | pow_eq_zero h' 83 | 84 | example (x y : ℝ) : x ^ 2 + y ^ 2 = 0 ↔ x = 0 ∧ y = 0 := by 85 | constructor 86 | · intro h 87 | constructor 88 | · exact aux h 89 | rw [add_comm] at h 90 | exact aux h 91 | rintro ⟨rfl, rfl⟩ 92 | norm_num 93 | 94 | example (x y : ℝ) : x ^ 2 + y ^ 2 = 0 ↔ x = 0 ∧ y = 0 := by 95 | have : (0 : ℝ) ^ 2 = 0 := by norm_num 96 | duper [*, aux, add_comm, add_zero] 97 | 98 | section 99 | 100 | -- Not good for Duper 101 | example (x : ℝ) : |x + 3| < 5 → -8 < x ∧ x < 2 := by 102 | rw [abs_lt] 103 | intro h 104 | constructor <;> linarith 105 | 106 | example : 3 ∣ Nat.gcd 6 15 := by 107 | rw [Nat.dvd_gcd_iff] 108 | constructor <;> norm_num 109 | 110 | -- Interesting: this uses `decide` at the end 111 | example : 3 ∣ Nat.gcd 6 15 := by 112 | duper [Nat.dvd_gcd_iff] {portfolioInstance := 7} 113 | 114 | end 115 | 116 | theorem not_monotone_iff {f : ℝ → ℝ} : ¬Monotone f ↔ ∃ x y, x ≤ y ∧ f x > f y := by 117 | rw [Monotone] 118 | push_neg 119 | rfl 120 | 121 | example {f : ℝ → ℝ} : ¬Monotone f ↔ ∃ x y, x ≤ y ∧ f x > f y := by 122 | duper [Monotone, not_le] 123 | 124 | example : ¬Monotone fun x : ℝ ↦ -x := by 125 | rw [not_monotone_iff] 126 | use 0, 1 127 | norm_num 128 | 129 | example : ¬Monotone fun x : ℝ ↦ -x := by 130 | duper [Monotone, not_lt, le_of_lt, neg_le_neg_iff, zero_lt_one] 131 | 132 | section 133 | variable {α : Type*} [PartialOrder α] 134 | variable (a b : α) 135 | 136 | example : a < b ↔ a ≤ b ∧ a ≠ b := by 137 | rw [lt_iff_le_not_le] 138 | constructor 139 | · rintro ⟨h0, h1⟩ 140 | constructor 141 | · exact h0 142 | intro h2 143 | apply h1 144 | rw [h2] 145 | rintro ⟨h0, h1⟩ 146 | constructor 147 | · exact h0 148 | intro h2 149 | apply h1 150 | apply le_antisymm h0 h2 151 | 152 | example : a < b ↔ a ≤ b ∧ a ≠ b := by 153 | duper [lt_iff_le_not_le, le_antisymm] 154 | 155 | end 156 | 157 | section 158 | variable {α : Type*} [Preorder α] 159 | variable (a b c : α) 160 | 161 | example : ¬a < a := by 162 | duper [lt_iff_le_not_le] 163 | 164 | example : a < b → b < c → a < c := by 165 | duper [lt_iff_le_not_le, le_trans] 166 | 167 | end 168 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/Prime.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | import Mathlib.Data.Nat.Prime 3 | import Mathlib.Tactic.Linarith 4 | 5 | namespace Nat 6 | 7 | set_option auto.tptp.solver.name "zipperposition" 8 | set_option throwPortfolioErrors true 9 | 10 | #check Nat.prime_def_lt -- Reproving this theorem using duper: 11 | 12 | theorem prime_def_lt_DUPER {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m < p, m ∣ p → m = 1 := by 13 | have : 2 ≤ p → 0 < p := by intro; linarith 14 | duper [*, Nat.prime_def_lt'', Nat.le_of_dvd, Nat.lt_of_le_of_ne, Nat.lt_irrefl] {portfolioInstance := 1} 15 | 16 | #check Nat.prime_def_lt' -- Reproving this theorem using duper: 17 | 18 | theorem prime_def_lt'_DUPER {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ (m : ℕ), 2 ≤ m → m < p → ¬m ∣ p := by 19 | constructor 20 | · have : ¬ 2 ≤ 1 := by duper [Nat.two_le_iff] 21 | duper [*, prime_def_lt_DUPER] 22 | · have : (∀ (m : ℕ), 2 ≤ m → m < p → ¬m ∣ p) → ∀ m < p, m ∣ p → m = 1 23 | · duper [Nat.two_le_iff, Nat.not_eq_zero_of_lt, Nat.eq_zero_of_zero_dvd] 24 | duper [*, prime_def_lt_DUPER] 25 | 26 | #check prime_def_le_sqrt -- Reproving this theorem using duper: 27 | 28 | set_option reduceInstances true -- The last two duper calls only work when reduceInstances is set to true 29 | theorem prime_def_le_sqrt_DUPER' {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m, 2 ≤ m → m ≤ sqrt p → ¬m ∣ p := by 30 | constructor 31 | · have : ∀ m, 2 ≤ m → 1 < m := by intros; linarith 32 | duper [*, prime_def_lt', sqrt_lt_self, Nat.lt_of_le_of_lt] 33 | · intro h 34 | rw [prime_def_lt'] 35 | refine ⟨h.1, ?_⟩ 36 | have h₂ : ∀ m, 2 ≤ m → m < p → m ∣ p → 2 ≤ (p / m) := by 37 | duper [*, Nat.lt_irrefl, Nat.dvd_iff_div_mul_eq, Nat.mul_zero, Nat.one_mul, two_le_iff] {portfolioInstance := 1} 38 | duper [*, Nat.div_dvd_of_dvd, Nat.dvd_iff_div_mul_eq, le_sqrt, Nat.mul_le_mul_right, 39 | Nat.mul_le_mul_left, mul_comm, Nat.le_total] {portfolioInstance := 1} 40 | -------------------------------------------------------------------------------- /DuperOnMathlib/DuperOnMathlib/Prime2.lean: -------------------------------------------------------------------------------- 1 | import Duper.Tactic 2 | import Mathlib.Data.Nat.Prime 3 | 4 | set_option includeExpensiveRules false 5 | 6 | theorem prime_def_lt''_new {p : ℕ} : Nat.Prime p ↔ 2 ≤ p ∧ ∀ (m) (_ : m ∣ p), m = 1 ∨ m = p := by 7 | have h1 : (1 : Nat) < 2 := @one_lt_two Nat _ _ _ _ _ 8 | have h2 : ∀ {a b c : ℕ}, a < b → b ≤ c → a < c := @LT.lt.trans_le Nat _ 9 | have h3 : ∀ {a b c : ℕ}, a ≠ 0 → (a * b = a * c ↔ b = c) := @mul_right_inj' Nat _ 10 | have h4 : ∀ {n m : ℕ}, n < m → 0 < m:= @pos_of_gt Nat _ 11 | have h5 : ∀ (a b : ℕ), a ∣ a * b := @dvd_mul_right Nat _ 12 | have h6 : ∀ (a : ℕ), a * 1 = a := @mul_one Nat _ 13 | have h7 : ∀ {x y : ℕ}, x < y → y ≠ x := @LT.lt.ne' Nat _ 14 | have h8 : ∀ {n : ℕ}, IsUnit n ↔ n = 1 := @Nat.isUnit_iff 15 | have h9 : ∀ {p : ℕ}, Nat.Prime p → 2 ≤ p := @Nat.Prime.two_le 16 | have h10 : ∀ {p : ℕ}, Nat.Prime p → ∀ (m : ℕ), m ∣ p → m = 1 ∨ m = p := @Nat.Prime.eq_one_or_self_of_dvd 17 | have h11 := Nat.irreducible_iff_nat_prime 18 | have h12 : ∀ {p : ℕ}, Irreducible p ↔ ¬IsUnit p ∧ ∀ (a b : ℕ), p = a * b → IsUnit a ∨ IsUnit b := @irreducible_iff Nat _ 19 | duper [*] {portfolioInstance := 0} 20 | 21 | set_option reduceInstances false 22 | 23 | -- Previously, turning reduceInstances off was sufficient to solve this problem. 24 | -- TODO: Look into why duper no longer solves it 25 | theorem prime_def_lt''_new2 {p : ℕ} : Nat.Prime p ↔ 2 ≤ p ∧ ∀ (m) (_ : m ∣ p), m = 1 ∨ m = p := by 26 | have h1 : (1 : Nat) < 2 := @one_lt_two Nat _ _ _ _ _ 27 | have h2 : ∀ {a b c : ℕ}, a < b → b ≤ c → a < c := @LT.lt.trans_le Nat _ 28 | have h3 : ∀ {a b c : ℕ}, a ≠ 0 → (a * b = a * c ↔ b = c) := @mul_right_inj' Nat _ 29 | have h4 : ∀ {n m : ℕ}, n < m → 0 < m:= @pos_of_gt Nat _ 30 | have h5 : ∀ (a b : ℕ), a ∣ a * b := @dvd_mul_right Nat _ 31 | have h6 : ∀ (a : ℕ), a * 1 = a := @mul_one Nat _ 32 | have h7 : ∀ {x y : ℕ}, x < y → y ≠ x := @LT.lt.ne' Nat _ 33 | have h8 : ∀ {n : ℕ}, IsUnit n ↔ n = 1 := @Nat.isUnit_iff 34 | have h9 : ∀ {p : ℕ}, Nat.Prime p → 2 ≤ p := @Nat.Prime.two_le 35 | have h10 : ∀ {p : ℕ}, Nat.Prime p → ∀ (m : ℕ), m ∣ p → m = 1 ∨ m = p := @Nat.Prime.eq_one_or_self_of_dvd 36 | have h11 := Nat.irreducible_iff_nat_prime 37 | have h12 : ∀ {p : ℕ}, Irreducible p ↔ ¬IsUnit p ∧ ∀ (a b : ℕ), p = a * b → IsUnit a ∨ IsUnit b := @irreducible_iff Nat _ 38 | duper [*] {portfolioInstance := 0} 39 | -------------------------------------------------------------------------------- /DuperOnMathlib/lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": 7, 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/leanprover/std4", 5 | "type": "git", 6 | "subDir": null, 7 | "rev": "9e37a01f8590f81ace095b56710db694b5bf8ca0", 8 | "name": "std", 9 | "manifestFile": "lake-manifest.json", 10 | "inputRev": "main", 11 | "inherited": true, 12 | "configFile": "lakefile.lean"}, 13 | {"url": "https://github.com/leanprover-community/quote4", 14 | "type": "git", 15 | "subDir": null, 16 | "rev": "d3a1d25f3eba0d93a58d5d3d027ffa78ece07755", 17 | "name": "Qq", 18 | "manifestFile": "lake-manifest.json", 19 | "inputRev": "master", 20 | "inherited": true, 21 | "configFile": "lakefile.lean"}, 22 | {"url": "https://github.com/leanprover-community/aesop", 23 | "type": "git", 24 | "subDir": null, 25 | "rev": "c7cff4551258d31c0d2d453b3f9cbca757d445f1", 26 | "name": "aesop", 27 | "manifestFile": "lake-manifest.json", 28 | "inputRev": "master", 29 | "inherited": true, 30 | "configFile": "lakefile.lean"}, 31 | {"url": "https://github.com/leanprover-community/ProofWidgets4", 32 | "type": "git", 33 | "subDir": null, 34 | "rev": "909febc72b4f64628f8d35cd0554f8a90b6e0749", 35 | "name": "proofwidgets", 36 | "manifestFile": "lake-manifest.json", 37 | "inputRev": "v0.0.23", 38 | "inherited": true, 39 | "configFile": "lakefile.lean"}, 40 | {"url": "https://github.com/leanprover/lean4-cli", 41 | "type": "git", 42 | "subDir": null, 43 | "rev": "a751d21d4b68c999accb6fc5d960538af26ad5ec", 44 | "name": "Cli", 45 | "manifestFile": "lake-manifest.json", 46 | "inputRev": "main", 47 | "inherited": true, 48 | "configFile": "lakefile.lean"}, 49 | {"url": "https://github.com/leanprover-community/mathlib4", 50 | "type": "git", 51 | "subDir": null, 52 | "rev": "75641abe702d35c946b6e59b8dc672c7e6d80b13", 53 | "name": "mathlib", 54 | "manifestFile": "lake-manifest.json", 55 | "inputRev": "v4.4.0-rc1", 56 | "inherited": false, 57 | "configFile": "lakefile.lean"}, 58 | {"url": "https://github.com/leanprover-community/lean-auto.git", 59 | "type": "git", 60 | "subDir": null, 61 | "rev": "b61345f6d42675298f8fa2b0cc5c471c264a6b40", 62 | "name": "auto", 63 | "manifestFile": "lake-manifest.json", 64 | "inputRev": "v0.0.4", 65 | "inherited": true, 66 | "configFile": "lakefile.lean"}, 67 | {"type": "path", 68 | "name": "Duper", 69 | "manifestFile": "lake-manifest.json", 70 | "inherited": false, 71 | "dir": "./../", 72 | "configFile": "lakefile.lean"}], 73 | "name": "duperOnMathlib", 74 | "lakeDir": ".lake"} 75 | -------------------------------------------------------------------------------- /DuperOnMathlib/lakefile.lean: -------------------------------------------------------------------------------- 1 | import Lake 2 | open Lake DSL 3 | 4 | package «duperOnMathlib» { 5 | -- add any package configuration options here 6 | } 7 | 8 | require «mathlib» from git "https://github.com/leanprover-community/mathlib4" @ "v4.4.0-rc1" 9 | 10 | require «Duper» from "../" 11 | 12 | @[default_target] 13 | lean_lib «DuperOnMathlib» { 14 | -- add library configuration options here 15 | } 16 | -------------------------------------------------------------------------------- /DuperOnMathlib/lakefile.olean: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/leanprover-community/duper/0101747d06ada379e28f73e23f3c39c5b2e40f55/DuperOnMathlib/lakefile.olean -------------------------------------------------------------------------------- /DuperOnMathlib/lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.4.0-rc1 2 | -------------------------------------------------------------------------------- /TODO.md: -------------------------------------------------------------------------------- 1 | # TODO 2 | 3 | I’ve listed some action items that could be taken to improve Duper. Most of these were previously 4 | listed in the earlier version of TODO.md, though I have removed some, changed some, and added new ones. 5 | 6 | - Better definition support (currently, we add the definition as a lemma, but we don’t enforce 7 | that that the definition ought to be unfolded). There are several examples in which the current 8 | handling of definitions appears to significantly hinder Duper’s ability to solve a problem. 9 | - Improving Duper’s inhabitation reasoning. Several of Duper’s current known bugs and limitations 10 | pertain to inhabitation reasoning, so expanding and/or restructuring Duper’s inhabitation reasoning 11 | could enable Duper to solve a variety of problems that it currently can not 12 | - Making curated collections of theorems tailored to particular problem domains. Examples 13 | include continuity lemmas, category theory lemmas, group lemmas, etc. 14 | - Finishing implementing inference rules: 15 | - BoolRw 16 | - ForAllRw 17 | - ExistsRw 18 | - Adding additional simplification rules: 19 | - Semantic tautology deletion? 20 | - Equality subsumption? 21 | - Infrastructure changes to better handle dependent terms: 22 | - For superposition, it might still make sense to test early whether a subterm is **depended** by 23 | another term because then we don't even need to add it to our term index. That could save us quite a 24 | bit of computation. 25 | - The current implementation of ```Lit.map``` and ```Lit.mapM``` is questionable. If $t : \alpha$ 26 | (at term level) and $f : Expr \to Expr$ (at meta level) is a function, then we probably do not always 27 | have $f \ t : f \ \alpha$. Currently it doesn't cause problem, but I expect there to be problems if we 28 | add support for dependent types. I think we should probably only map to ```lhs``` and ```rhs```, and 29 | then update the type using ```Meta.inferType```. 30 | - Improve indexing functions for fingerprint indices 31 | - Current fingerprint function is from http://www.eprover.eu/EXPDATA/FP_INDEX/schulz_fp-index_ext.pdf. 32 | Explore whether other fingerprint functions would be better (especially those with more features) 33 | - Look into whether superposition and demodulation are taking an excessive amount of time eliminating bad 34 | potential partners. If so, we might be able to save some time by checking the type of potential partners 35 | sooner (currently, we don't check until Unif.lean/Match.lean). The reason this might be helpful is that 36 | whenever I look for potential unification targets for a metavariable, I'm finding *everything* in the index. 37 | This is appropriate and necessary in FOL, but for our setting, the types should be able to rule most of these 38 | bad potential partners out. 39 | - Even better than checking the types earlier in superposition/demodulation would be to check them in the 40 | indexing structures. If we indexed expressions not just to their Clause X ClausePos pair, but to a 41 | Clause X ClausePos X Expr type (with the final Expr being the type of the indexed subexpression), then when 42 | we search for e's unification partners, we can easily return only those whose type is compatible with e's. 43 | - Alternatively, we could add a top-level component to RootCFPTrie so that rather than just having one root 44 | ClauseFingerprintTrie, we have a HashMap from types to ClauseFingerprintTries (where, if type tau maps to 45 | ClauseFingerprintTrie t, then t only indexes expressions of type tau). I don't think this would significantly 46 | increase memory, and depending on how much time is being wasted eliminating bad potential partners in 47 | superposition.lean and demodulation.lean, it might save a lot of time. 48 | - Because the types of SimpRule and BackwardSimpRule were changed to take Clauses rather than MClauses as input, 49 | calling loadClause has been moved from Simp.lean to the beginning of each simplification rule. This isn't a 50 | problem, but there are some places where it may not be necessary to call loadClause. Going through the 51 | simplification rules and changing them to only call loadClause if it is actually required may make some of the 52 | simplification rules more efficient 53 | - It is presently unclear to me whether this improvement would be significant, because it is has the potential 54 | to remove many calls to loadClause, or negligible because loadClause is implemented efficiently 55 | - Use mvars in Clause to avoid cost of constantly converting between the Clause and MClause representations? 56 | - Update Expr.weight function to be more similar to Zipperposition's ho_weight function 57 | - Currently, we have a hacky implementation of removing clauses from fingerprint indexes (tacking on a filter 58 | before retrieving). If this turns out to be too inefficient, implement removal from fingerprint indexes properly. 59 | - Add inter-simplification to more faithfully implement immediate simplification 60 | - Investigate why the current implementation of orphan elimination (removeDescendants) worsens performance on average -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/leanprover-community/batteries", 5 | "type": "git", 6 | "subDir": null, 7 | "scope": "", 8 | "rev": "cb53cd3f2958ee503f97ea6162929cabaf96bde9", 9 | "name": "batteries", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": "v4.20.0-rc2", 12 | "inherited": false, 13 | "configFile": "lakefile.toml"}, 14 | {"url": "https://github.com/leanprover-community/lean-auto.git", 15 | "type": "git", 16 | "subDir": null, 17 | "scope": "", 18 | "rev": "b9a5fa7a88410474c13905bd23d331ea079923b3", 19 | "name": "auto", 20 | "manifestFile": "lake-manifest.json", 21 | "inputRev": "b9a5fa7a88410474c13905bd23d331ea079923b3", 22 | "inherited": false, 23 | "configFile": "lakefile.lean"}], 24 | "name": "Duper", 25 | "lakeDir": ".lake"} 26 | -------------------------------------------------------------------------------- /lakefile.lean: -------------------------------------------------------------------------------- 1 | import Lake 2 | 3 | open Lake DSL 4 | 5 | require auto from git "https://github.com/leanprover-community/lean-auto.git"@"b9a5fa7a88410474c13905bd23d331ea079923b3" 6 | require batteries from git "https://github.com/leanprover-community/batteries" @ "v4.20.0-rc2" 7 | 8 | package Duper { 9 | precompileModules := true 10 | preferReleaseBuild := false 11 | } 12 | 13 | lean_lib Duper 14 | 15 | @[default_target] 16 | lean_exe duper { 17 | root := `Main 18 | } 19 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.20.0-rc5 2 | -------------------------------------------------------------------------------- /starexec/.gitignore: -------------------------------------------------------------------------------- 1 | bin/duper 2 | duper_starexec.tar -------------------------------------------------------------------------------- /starexec/bin/starexec_run_default.sh: -------------------------------------------------------------------------------- 1 | #!/bin/sh 2 | 3 | ./duper "$1" -------------------------------------------------------------------------------- /starexec/pack_starexec.sh: -------------------------------------------------------------------------------- 1 | #!/bin/bash 2 | parent_path=$( cd "$(dirname "${BASH_SOURCE[0]}")" ; pwd -P ) 3 | 4 | cp $parent_path/../build/bin/duper $parent_path/bin/duper 5 | cd $parent_path 6 | tar -cf duper_starexec.tar bin --------------------------------------------------------------------------------