├── .github └── workflows │ ├── update_versions.yml │ └── upgrade.yml ├── .gitignore ├── .travis.yml ├── README.md ├── leanpkg.toml ├── src └── super │ ├── cdcl.lean │ ├── cdcl_solver.lean │ ├── clause.lean │ ├── clause_ops.lean │ ├── clausifier.lean │ ├── datatypes.lean │ ├── default.lean │ ├── defs.lean │ ├── demod.lean │ ├── equality.lean │ ├── factoring.lean │ ├── inhabited.lean │ ├── lpo.lean │ ├── misc_preprocessing.lean │ ├── prover.lean │ ├── prover_state.lean │ ├── resolution.lean │ ├── selection.lean │ ├── simp.lean │ ├── splitting.lean │ ├── subsumption.lean │ ├── superposition.lean │ ├── trim.lean │ └── utils.lean └── test ├── cdcl_examples.lean ├── super_examples.lean └── super_tests.lean /.github/workflows/update_versions.yml: -------------------------------------------------------------------------------- 1 | on: 2 | push: 3 | branches: 4 | - master 5 | 6 | jobs: 7 | update_lean_xyz_branch: 8 | runs-on: ubuntu-latest 9 | name: Update lean-x.y.z branch 10 | steps: 11 | 12 | - name: checkout project 13 | uses: actions/checkout@v2 14 | with: 15 | fetch-depth: 0 16 | 17 | - name: update branch 18 | uses: leanprover-contrib/update-versions-action@master -------------------------------------------------------------------------------- /.github/workflows/upgrade.yml: -------------------------------------------------------------------------------- 1 | on: 2 | schedule: 3 | - cron: '0 0 * * *' 4 | 5 | jobs: 6 | upgrade_lean: 7 | runs-on: ubuntu-latest 8 | name: Bump Lean and dependency versions 9 | steps: 10 | - name: checkout project 11 | uses: actions/checkout@v2 12 | - name: upgrade Lean and dependencies 13 | uses: leanprover-contrib/lean-upgrade-action@master 14 | with: 15 | repo: ${{ github.repository }} 16 | access-token: ${{ secrets.GITHUB_TOKEN }} 17 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | *.olean 2 | /_target 3 | /leanpkg.path 4 | -------------------------------------------------------------------------------- /.travis.yml: -------------------------------------------------------------------------------- 1 | language: generic 2 | 3 | install: 4 | - curl https://raw.githubusercontent.com/Kha/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y 5 | - source ~/.elan/env 6 | 7 | script: 8 | - leanpkg test 9 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # super 2 | 3 | [![Build Status](https://travis-ci.org/leanprover/super.svg?branch=master)](https://travis-ci.org/leanprover/super) 4 | 5 | A super position prover for Lean. 6 | 7 | TODO: add docs 8 | -------------------------------------------------------------------------------- /leanpkg.toml: -------------------------------------------------------------------------------- 1 | [package] 2 | name = "super" 3 | version = "0.1" 4 | lean_version = "leanprover-community/lean:3.50.3" 5 | path = "src" 6 | 7 | [dependencies] 8 | -------------------------------------------------------------------------------- /src/super/cdcl.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause .clausifier .cdcl_solver 7 | open tactic expr monad super 8 | 9 | meta def cdcl_t (th_solver : tactic unit) : tactic unit := do 10 | as_refutation, local_false ← target, 11 | clauses ← clauses_of_context, clauses ← get_clauses_classical clauses, 12 | clauses.mmap' (λc, do c_pp ← pp c, clause.validate c <|> fail c_pp), 13 | res ← cdcl.solve (cdcl.theory_solver_of_tactic th_solver) local_false clauses, 14 | match res with 15 | | (cdcl.result.unsat proof) := exact proof 16 | | (cdcl.result.sat interp) := 17 | let interp' := do e ← interp.to_list, [if e.2 = tt then e.1 else `(not %%e.1)] in 18 | do pp_interp ← pp interp', 19 | fail (to_fmt "satisfying assignment: " ++ pp_interp) 20 | end 21 | 22 | meta def cdcl : tactic unit := cdcl_t skip 23 | -------------------------------------------------------------------------------- /src/super/cdcl_solver.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause 7 | open tactic expr monad super native 8 | 9 | namespace cdcl 10 | 11 | @[reducible] meta def prop_var := expr 12 | @[reducible] meta def proof_term := expr 13 | @[reducible] meta def proof_hyp := expr 14 | 15 | meta inductive trail_elem 16 | | dec : prop_var → bool → proof_hyp → trail_elem 17 | | propg : prop_var → bool → proof_term → proof_hyp → trail_elem 18 | | dbl_neg_propg : prop_var → bool → proof_term → proof_hyp → trail_elem 19 | 20 | namespace trail_elem 21 | 22 | meta def var : trail_elem → prop_var 23 | | (dec v _ _) := v 24 | | (propg v _ _ _) := v 25 | | (dbl_neg_propg v _ _ _) := v 26 | 27 | meta def phase : trail_elem → bool 28 | | (dec _ ph _) := ph 29 | | (propg _ ph _ _) := ph 30 | | (dbl_neg_propg _ ph _ _) := ph 31 | 32 | meta def hyp : trail_elem → proof_hyp 33 | | (dec _ _ h) := h 34 | | (propg _ _ _ h) := h 35 | | (dbl_neg_propg _ _ _ h) := h 36 | 37 | meta def is_decision : trail_elem → bool 38 | | (dec _ _ _) := tt 39 | | (propg _ _ _ _) := ff 40 | | (dbl_neg_propg _ _ _ _) := ff 41 | 42 | end trail_elem 43 | 44 | meta structure var_state := 45 | (phase : bool) 46 | (assigned : option proof_hyp) 47 | 48 | meta structure learned_clause := 49 | (c : clause) 50 | (actual_proof : proof_term) 51 | 52 | meta inductive prop_lit 53 | | neg : prop_var → prop_lit 54 | | pos : prop_var → prop_lit 55 | 56 | namespace prop_lit 57 | 58 | meta instance prop_lit.has_lt : has_lt prop_lit := 59 | ⟨λl₁ l₂, match l₁, l₂ with 60 | | pos _, neg _ := true 61 | | neg _, pos _ := false 62 | | pos v₁, pos v₂ := v₁ < v₂ 63 | | neg v₁, neg v₂ := v₁ < v₂ 64 | end⟩ 65 | 66 | meta instance prop_lit.decidable_lt : decidable_rel ((<) : prop_lit → prop_lit → Prop) := 67 | λl₁ l₂, match l₁, l₂ with 68 | | pos _, neg _ := is_true trivial 69 | | neg _, pos _ := is_false not_false 70 | | pos v₁, pos v₂ := expr.lt_prop.decidable_rel v₁ v₂ 71 | | neg v₁, neg v₂ := expr.lt_prop.decidable_rel v₁ v₂ 72 | end 73 | 74 | meta def of_cls_lit : clause.literal → prop_lit 75 | | (clause.literal.left v) := neg v 76 | | (clause.literal.right v) := pos v 77 | 78 | meta def of_var_and_phase (v : prop_var) : bool → prop_lit 79 | | tt := pos v 80 | | ff := neg v 81 | 82 | end prop_lit 83 | 84 | meta def watch_map := rb_map name (ℕ × ℕ × clause) 85 | 86 | meta structure state := 87 | (trail : list trail_elem) 88 | (vars : rb_map prop_var var_state) 89 | (unassigned : rb_map prop_var prop_var) 90 | (clauses : list clause) 91 | (learned : list learned_clause) 92 | (watches : rb_map prop_lit watch_map) 93 | (conflict : option proof_term) 94 | (unitp_queue : list prop_var) 95 | (local_false : expr) 96 | 97 | namespace state 98 | 99 | meta def initial (local_false : expr) : state := { 100 | trail := [], 101 | vars := rb_map.mk _ _, 102 | unassigned := rb_map.mk _ _, 103 | clauses := [], 104 | learned := [], 105 | watches := rb_map.mk _ _, 106 | conflict := none, 107 | unitp_queue := [], 108 | local_false := local_false 109 | } 110 | 111 | meta def watches_for (st : state) (pl : prop_lit) : watch_map := 112 | (st.watches.find pl).get_or_else (rb_map.mk _ _) 113 | 114 | end state 115 | 116 | meta def solver := state_t state tactic 117 | section 118 | local attribute [reducible] solver 119 | meta instance : monad solver := infer_instance 120 | meta instance : alternative solver := infer_instance 121 | meta instance : monad_state _ solver := infer_instance 122 | meta instance : has_monad_lift tactic solver := infer_instance 123 | meta instance (α : Type) : has_coe (tactic α) (solver α) := 124 | ⟨monad_lift⟩ 125 | end 126 | 127 | meta def fail {A B} [has_to_format B] (b : B) : solver A := 128 | @tactic.fail A B _ b 129 | 130 | meta def get_local_false : solver expr := 131 | do st ← get, return st.local_false 132 | 133 | meta def mk_var_core (v : prop_var) (ph : bool) : solver unit := do 134 | modify $ λst, match st.vars.find v with 135 | | (some _) := st 136 | | none := { st with 137 | vars := st.vars.insert v ⟨ph, none⟩, 138 | unassigned := st.unassigned.insert v v 139 | } 140 | end, 141 | skip 142 | 143 | meta def mk_var (v : prop_var) : solver unit := mk_var_core v ff 144 | 145 | meta def set_conflict (proof : proof_term) : solver unit := do 146 | modify $ λst, { st with conflict := some proof }, skip 147 | 148 | meta def has_conflict : solver bool := 149 | do st ← get, return st.conflict.is_some 150 | 151 | meta def push_trail (elem : trail_elem) : solver unit := do 152 | st ← get, 153 | match st.vars.find elem.var with 154 | | none := fail $ "unknown variable: " ++ elem.var.to_string 155 | | some ⟨_, some _⟩ := fail $ "adding already assigned variable to trail: " ++ elem.var.to_string 156 | | some ⟨_, none⟩ := 157 | put { st with 158 | vars := st.vars.insert elem.var ⟨elem.phase, some elem.hyp⟩, 159 | unassigned := st.unassigned.erase elem.var, 160 | trail := elem :: st.trail, 161 | unitp_queue := elem.var :: st.unitp_queue 162 | } >> skip 163 | end 164 | 165 | meta def pop_trail_core : solver (option trail_elem) := do 166 | st ← get, 167 | match st.trail with 168 | | elem :: rest := do 169 | put { st with trail := rest, 170 | vars := st.vars.insert elem.var ⟨elem.phase, none⟩, 171 | unassigned := st.unassigned.insert elem.var elem.var, 172 | unitp_queue := [] }, 173 | return $ some elem 174 | | [] := return none 175 | end 176 | 177 | meta def is_decision_level_zero : solver bool := 178 | do st ← get, return $ st.trail.for_all $ λelem, ¬elem.is_decision 179 | 180 | meta def revert_to_decision_level_zero : unit → solver unit | () := do 181 | is_dl0 ← is_decision_level_zero, 182 | if is_dl0 then return () 183 | else do pop_trail_core, revert_to_decision_level_zero () 184 | 185 | meta def formula_of_lit (local_false : expr) (v : prop_var) (ph : bool) := 186 | if ph then v else imp v local_false 187 | 188 | meta def lookup_var (v : prop_var) : solver (option var_state) := 189 | do st ← get, return $ st.vars.find v 190 | 191 | meta def add_propagation (v : prop_var) (ph : bool) (just : proof_term) (just_is_dn : bool) : solver unit := 192 | do v_st ← lookup_var v, local_false ← get_local_false, match v_st with 193 | | none := fail $ "propagating unknown variable: " ++ v.to_string 194 | | some ⟨assg_ph, some proof⟩ := 195 | if ph = assg_ph then 196 | return () 197 | else if assg_ph ∧ ¬just_is_dn then 198 | set_conflict (app just proof) 199 | else 200 | set_conflict (app proof just) 201 | | some ⟨_, none⟩ := do 202 | hyp_name ← mk_fresh_name, 203 | hyp ← return $ local_const hyp_name hyp_name binder_info.default (formula_of_lit local_false v ph), 204 | if just_is_dn then do 205 | push_trail $ trail_elem.dbl_neg_propg v ph just hyp 206 | else do 207 | push_trail $ trail_elem.propg v ph just hyp 208 | end 209 | 210 | meta def add_decision (v : prop_var) (ph : bool) : solver unit := do 211 | hyp_name ← mk_fresh_name, 212 | local_false ← get_local_false, 213 | hyp ← return $ local_const hyp_name hyp_name binder_info.default (formula_of_lit local_false v ph), 214 | push_trail $ trail_elem.dec v ph hyp 215 | 216 | meta def lookup_lit (l : clause.literal) : solver (option (bool × proof_hyp)) := 217 | do var_st_opt ← lookup_var l.formula, match var_st_opt with 218 | | none := return none 219 | | some ⟨ph, none⟩ := return none 220 | | some ⟨ph, some proof⟩ := 221 | return $ some (if l.is_neg then bnot ph else ph, proof) 222 | end 223 | 224 | meta def lit_is_false (l : clause.literal) : solver bool := 225 | do s ← lookup_lit l, return $ match s with 226 | | some (ff, _) := tt 227 | | _ := ff 228 | end 229 | 230 | meta def lit_is_not_false (l : clause.literal) : solver bool := 231 | do isf ← lit_is_false l, return $ bnot isf 232 | 233 | meta def cls_is_false (c : clause) : solver bool := 234 | list.band <$> mapm lit_is_false c.get_lits 235 | 236 | private meta def unit_propg_cls' : clause → solver (option prop_var) | c := 237 | if c.num_lits = 0 then return (some c.proof) 238 | else let hd := c.get_lit 0 in 239 | do lit_st ← lookup_lit hd, match lit_st with 240 | | some (ff, isf_prf) := unit_propg_cls' (c.inst isf_prf) 241 | | _ := return none 242 | end 243 | 244 | meta def unit_propg_cls : clause → solver unit | c := 245 | do has_confl ← has_conflict, 246 | if has_confl then return () else 247 | if c.num_lits = 0 then do set_conflict c.proof 248 | else let hd := c.get_lit 0 in 249 | do lit_st ← lookup_lit hd, match lit_st with 250 | | some (ff, isf_prf) := unit_propg_cls (c.inst isf_prf) 251 | | some (tt, _) := return () 252 | | none := 253 | do fls_prf_opt ← unit_propg_cls' (c.inst (expr.mk_var 0)), 254 | match fls_prf_opt with 255 | | some fls_prf := do 256 | fls_prf' ← return $ lam `H binder_info.default c.type.binding_domain fls_prf, 257 | if hd.is_neg then 258 | add_propagation hd.formula ff fls_prf' ff 259 | else 260 | add_propagation hd.formula tt fls_prf' tt 261 | | none := return () 262 | end 263 | end 264 | 265 | private meta def modify_watches_for (pl : prop_lit) (f : watch_map → watch_map) : solver unit := do 266 | modify $ λst, { st with 267 | watches := st.watches.insert pl $ f $ st.watches_for pl 268 | }, skip 269 | 270 | private meta def add_watch (n : name) (c : clause) (i j : ℕ) : solver unit := 271 | let l := c.get_lit i, pl := prop_lit.of_cls_lit l in 272 | modify_watches_for pl $ λw, w.insert n (i,j,c) 273 | 274 | private meta def remove_watch (n : name) (c : clause) (i : ℕ) : solver unit := 275 | let l := c.get_lit i, pl := prop_lit.of_cls_lit l in 276 | modify_watches_for pl $ λw, w.erase n 277 | 278 | private meta def set_watches (n : name) (c : clause) : solver unit := 279 | if c.num_lits = 0 then 280 | set_conflict c.proof 281 | else if c.num_lits = 1 then 282 | unit_propg_cls c 283 | else do 284 | not_false_lits ← filter (λi, lit_is_not_false (c.get_lit i)) (list.range c.num_lits), 285 | match not_false_lits with 286 | | [] := do 287 | add_watch n c 0 1, 288 | add_watch n c 1 0, 289 | unit_propg_cls c 290 | | [i] := 291 | let j := if i = 0 then 1 else 0 in do 292 | add_watch n c i j, 293 | add_watch n c j i, 294 | unit_propg_cls c 295 | | (i::j::_) := do 296 | add_watch n c i j, 297 | add_watch n c j i 298 | end 299 | 300 | meta def update_watches (n : name) (c : clause) (i₁ i₂ : ℕ) : solver unit := do 301 | remove_watch n c i₁, 302 | remove_watch n c i₂, 303 | set_watches n c 304 | 305 | meta def mk_clause (c : clause) : solver unit := do 306 | c : clause ← c.distinct, 307 | c.get_lits.mmap' (λl, mk_var l.formula), 308 | revert_to_decision_level_zero (), 309 | modify $ λst, { st with clauses := c :: st.clauses }, 310 | c_name ← mk_fresh_name, 311 | set_watches c_name c 312 | 313 | meta def unit_propg_var (v : prop_var) : solver unit := 314 | do st ← get, if st.conflict.is_some then return () else 315 | match st.vars.find v with 316 | | some ⟨ph, none⟩ := fail $ "propagating unassigned variable: " ++ v.to_string 317 | | none := fail $ "unknown variable: " ++ v.to_string 318 | | some ⟨ph, some _⟩ := 319 | let watches := st.watches_for $ prop_lit.of_var_and_phase v (bnot ph) in 320 | watches.to_list.mmap' $ λw, update_watches w.1 w.2.2.2 w.2.1 w.2.2.1 321 | end 322 | 323 | meta def analyze_conflict' (local_false : expr) : proof_term → list trail_elem → clause 324 | | proof (trail_elem.dec v ph hyp :: es) := 325 | let abs_prf := abstract_local proof hyp.local_uniq_name in 326 | if has_var abs_prf then 327 | clause.close_const (analyze_conflict' proof es) hyp 328 | else 329 | analyze_conflict' proof es 330 | | proof (trail_elem.propg v ph l_prf hyp :: es) := 331 | let abs_prf := abstract_local proof hyp.local_uniq_name in 332 | if has_var abs_prf then 333 | analyze_conflict' (app (lam hyp.local_pp_name binder_info.default (formula_of_lit local_false v ph) abs_prf) l_prf) es 334 | else 335 | analyze_conflict' proof es 336 | | proof (trail_elem.dbl_neg_propg v ph l_prf hyp :: es) := 337 | let abs_prf := abstract_local proof hyp.local_uniq_name in 338 | if has_var abs_prf then 339 | analyze_conflict' (app l_prf (lambdas [hyp] proof)) es 340 | else 341 | analyze_conflict' proof es 342 | | proof [] := ⟨0, 0, proof, local_false, local_false⟩ 343 | 344 | meta def analyze_conflict (proof : proof_term) : solver clause := 345 | do st ← get, return $ analyze_conflict' st.local_false proof st.trail 346 | 347 | meta def add_learned (c : clause) : solver unit := do 348 | prf_abbrev_name ← mk_fresh_name, 349 | c' ← return { c with proof := local_const prf_abbrev_name prf_abbrev_name binder_info.default c.type }, 350 | modify $ λst, { st with learned := ⟨c', c.proof⟩ :: st.learned }, 351 | c_name ← mk_fresh_name, 352 | set_watches c_name c' 353 | 354 | meta def backtrack_with : clause → solver unit | conflict_clause := do 355 | isf ← cls_is_false conflict_clause, 356 | if ¬isf then 357 | modify (λst, { st with conflict := none }) >> skip 358 | else do 359 | removed_elem ← pop_trail_core, 360 | if removed_elem.is_some then 361 | backtrack_with conflict_clause 362 | else 363 | return () 364 | 365 | meta def replace_learned_clauses' : proof_term → list learned_clause → proof_term 366 | | proof [] := proof 367 | | proof (⟨c, actual_proof⟩ :: lcs) := 368 | let abs_prf := abstract_local proof c.proof.local_uniq_name in 369 | if has_var abs_prf then 370 | replace_learned_clauses' (elet c.proof.local_pp_name c.type actual_proof abs_prf) lcs 371 | else 372 | replace_learned_clauses' proof lcs 373 | 374 | meta def replace_learned_clauses (proof : proof_term) : solver proof_term := 375 | do st ← get, return $ replace_learned_clauses' proof st.learned 376 | 377 | meta inductive result 378 | | unsat : proof_term → result 379 | | sat : rb_map prop_var bool → result 380 | 381 | variable theory_solver : solver (option proof_term) 382 | 383 | meta def unit_propg : unit → solver unit | () := do 384 | st ← get, 385 | if st.conflict.is_some then return () else 386 | match st.unitp_queue with 387 | | [] := return () 388 | | (v::vs) := do 389 | put { st with unitp_queue := vs }, 390 | unit_propg_var v, 391 | unit_propg () 392 | end 393 | 394 | private meta def run' : unit → solver result | () := do 395 | unit_propg (), 396 | st ← get, 397 | match st.conflict with 398 | | some conflict := do 399 | conflict_clause ← analyze_conflict conflict, 400 | if conflict_clause.num_lits = 0 then do 401 | proof ← replace_learned_clauses conflict_clause.proof, 402 | return (result.unsat proof) 403 | else do 404 | backtrack_with conflict_clause, 405 | add_learned conflict_clause, 406 | run' () 407 | | none := 408 | match st.unassigned.min with 409 | | none := do 410 | theory_conflict ← theory_solver, 411 | match theory_conflict with 412 | | some conflict := do set_conflict conflict, run' () 413 | | none := return $ result.sat (st.vars.map (λvar_st, var_st.phase)) 414 | end 415 | | some unassigned := 416 | match st.vars.find unassigned with 417 | | some ⟨ph, none⟩ := do add_decision unassigned ph, run' () 418 | | _ := fail $ "unassigned variable is assigned: " ++ unassigned.to_string 419 | end 420 | end 421 | end 422 | 423 | meta def run : solver result := run' theory_solver () 424 | 425 | meta def solve (local_false : expr) (clauses : list clause) : tactic result := do 426 | res ← (do clauses.mmap' mk_clause, run theory_solver).run (state.initial local_false), 427 | return res.1 428 | 429 | meta def theory_solver_of_tactic (th_solver : tactic unit) : cdcl.solver (option cdcl.proof_term) := 430 | do s ← get, ↑do 431 | hyps ← return $ s.trail.map (λe, e.hyp), 432 | subgoal ← mk_meta_var s.local_false, 433 | goals ← get_goals, 434 | set_goals [subgoal], 435 | hvs ← hyps.mmap' (λhyp, assertv hyp.local_pp_name hyp.local_type hyp), 436 | solved ← (do th_solver, done, return tt) <|> return ff, 437 | set_goals goals, 438 | if solved then do 439 | proof ← instantiate_mvars subgoal, 440 | proof' ← whnf proof, -- gets rid of the unnecessary asserts 441 | return $ some proof' 442 | else 443 | return none 444 | 445 | end cdcl 446 | -------------------------------------------------------------------------------- /src/super/clause.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import init.meta.tactic .utils .trim 7 | open expr list tactic monad decidable native 8 | 9 | namespace super 10 | 11 | meta def is_local_not (local_false : expr) (e : expr) : option expr := 12 | match e with 13 | | (pi _ _ a b) := if b = local_false then some a else none 14 | | _ := if local_false = `(false) then is_not e else none 15 | end 16 | 17 | meta structure clause := 18 | (num_quants : ℕ) 19 | (num_lits : ℕ) 20 | (proof : expr) 21 | (type : expr) 22 | (local_false : expr) 23 | 24 | namespace clause 25 | 26 | meta def num_binders (c : clause) : ℕ := num_quants c + num_lits c 27 | 28 | meta def inst (c : clause) (e : expr) : clause := 29 | (if num_quants c > 0 30 | then mk (num_quants c - 1) (num_lits c) 31 | else mk 0 (num_lits c - 1)) 32 | (app (proof c) e) (instantiate_var (binding_body (type c)) e) c.local_false 33 | 34 | meta def instn (c : clause) (es : list expr) : clause := 35 | foldr (λe c', inst c' e) c es 36 | 37 | meta def open_const (c : clause) : tactic (clause × expr) := do 38 | n ← mk_fresh_name, 39 | b ← return $ local_const n (binding_name (type c)) (binding_info (type c)) (binding_domain (type c)), 40 | return (inst c b, b) 41 | 42 | meta def open_meta (c : clause) : tactic (clause × expr) := do 43 | b ← mk_meta_var (binding_domain (type c)), 44 | return (inst c b, b) 45 | 46 | meta def close_const (c : clause) (e : expr) : clause := 47 | match e with 48 | | local_const uniq pp info t := 49 | let abst_type' := abstract_local (type c) (local_uniq_name e) in 50 | let type' := pi pp binder_info.default t (abstract_local (type c) uniq) in 51 | let abs_prf := abstract_local (proof c) uniq in 52 | let proof' := lambdas [e] c.proof in 53 | if num_quants c > 0 ∨ has_var abst_type' then 54 | { c with num_quants := c.num_quants + 1, proof := proof', type := type' } 55 | else 56 | { c with num_lits := c.num_lits + 1, proof := proof', type := type' } 57 | | _ := ⟨0, 0, default, default, default⟩ 58 | end 59 | 60 | meta def open_constn : clause → ℕ → tactic (clause × list expr) 61 | | c 0 := return (c, nil) 62 | | c (n+1) := do 63 | (c', b) ← open_const c, 64 | (c'', bs) ← open_constn c' n, 65 | return (c'', b::bs) 66 | 67 | meta def open_metan : clause → ℕ → tactic (clause × list expr) 68 | | c 0 := return (c, nil) 69 | | c (n+1) := do 70 | (c', b) ← open_meta c, 71 | (c'', bs) ← open_metan c' n, 72 | return (c'', b::bs) 73 | 74 | meta def close_constn : clause → list expr → clause 75 | | c [] := c 76 | | c (b::bs') := close_const (close_constn c bs') b 77 | 78 | private meta def parse_clause (local_false : expr) : expr → expr → tactic clause 79 | | proof (pi n bi d b) := do 80 | lc_n ← mk_fresh_name, 81 | lc ← return $ local_const lc_n n bi d, 82 | c ← parse_clause (app proof lc) (instantiate_var b lc), 83 | return $ c.close_const $ local_const lc_n n binder_info.default d 84 | | proof `(not %%formula) := parse_clause proof (formula.imp `(false)) 85 | | proof type := 86 | if type = local_false then do 87 | return { num_quants := 0, num_lits := 0, proof := proof, type := type, local_false := local_false } 88 | else do 89 | univ ← infer_univ type, 90 | not_type ← return $ imp type local_false, 91 | parse_clause (lam `H binder_info.default not_type $ app (mk_var 0) proof) (imp not_type local_false) 92 | 93 | meta def of_proof_and_type (local_false proof type : expr) : tactic clause := 94 | parse_clause local_false proof type 95 | 96 | meta def of_proof (local_false proof : expr) : tactic clause := do 97 | type ← infer_type proof, 98 | of_proof_and_type local_false proof type 99 | 100 | meta def of_classical_proof (proof : expr) : tactic clause := 101 | of_proof `(false) proof 102 | 103 | meta def inst_mvars (c : clause) : tactic clause := do 104 | proof' ← instantiate_mvars (proof c), 105 | type' ← instantiate_mvars (type c), 106 | return { c with proof := proof', type := type' } 107 | 108 | meta inductive literal 109 | | left : expr → literal 110 | | right : expr → literal 111 | 112 | namespace literal 113 | 114 | meta instance : decidable_eq literal := by mk_dec_eq_instance 115 | 116 | meta def formula : literal → expr 117 | | (left f) := f 118 | | (right f) := f 119 | 120 | meta def is_neg : literal → bool 121 | | (left _) := tt 122 | | (right _) := ff 123 | 124 | meta def is_pos (l : literal) : bool := bnot l.is_neg 125 | 126 | meta def to_formula (l : literal) : expr := 127 | if l.is_neg 128 | then app (const ``not []) l.formula 129 | else formula l 130 | 131 | meta def type_str : literal → string 132 | | (literal.left _) := "left" 133 | | (literal.right _) := "right" 134 | 135 | meta instance : has_to_tactic_format literal := 136 | ⟨λl, do 137 | pp_f ← pp l.formula, 138 | return $ to_fmt l.type_str ++ " (" ++ pp_f ++ ")"⟩ 139 | 140 | end literal 141 | 142 | private meta def get_binding_body : expr → ℕ → expr 143 | | e 0 := e 144 | | e (i+1) := get_binding_body e.binding_body i 145 | 146 | meta def get_binder (e : expr) (i : nat) := 147 | binding_domain (get_binding_body e i) 148 | 149 | meta def validate (c : clause) : tactic unit := do 150 | concl ← return $ get_binding_body c.type c.num_binders, 151 | unify concl c.local_false 152 | <|> (do pp_concl ← pp concl, pp_lf ← pp c.local_false, 153 | fail $ to_fmt "wrong local false: " ++ pp_concl ++ " =!= " ++ pp_lf), 154 | type' ← infer_type c.proof, 155 | unify c.type type' <|> (do pp_ty ← pp c.type, pp_ty' ← pp type', 156 | fail (to_fmt "wrong type: " ++ pp_ty ++ " =!= " ++ pp_ty')) 157 | 158 | meta def get_lit (c : clause) (i : nat) : literal := 159 | let bind := get_binder (type c) (num_quants c + i) in 160 | match is_local_not c.local_false bind with 161 | | some formula := literal.right formula 162 | | none := literal.left bind 163 | end 164 | 165 | meta def lits_where (c : clause) (p : literal → bool) : list nat := 166 | list.filter (λl, p (get_lit c l)) (range (num_lits c)) 167 | 168 | meta def get_lits (c : clause) : list literal := 169 | list.map (get_lit c) (range c.num_lits) 170 | 171 | private meta def tactic_format (c : clause) : tactic format := do 172 | c ← c.open_metan c.num_quants, 173 | pp (do l ← c.1.get_lits, [l.to_formula]) 174 | 175 | meta instance : has_to_tactic_format clause := ⟨tactic_format⟩ 176 | 177 | meta def is_maximal (gt : expr → expr → bool) (c : clause) (i : nat) : bool := 178 | list.empty (list.filter (λj, gt (get_lit c j).formula (get_lit c i).formula) (range c.num_lits)) 179 | 180 | meta def normalize (c : clause) : tactic clause := do 181 | opened ← open_constn c (num_binders c), 182 | lconsts_in_types ← return $ contained_lconsts_list (list.map local_type opened.2), 183 | quants' ← return $ list.filter (λlc, rb_map.contains lconsts_in_types (local_uniq_name lc)) 184 | opened.2, 185 | lits' ← return $ list.filter (λlc, ¬rb_map.contains lconsts_in_types (local_uniq_name lc)) 186 | opened.2, 187 | return $ close_constn opened.1 (quants' ++ lits') 188 | 189 | meta def whnf_head_lit (c : clause) : tactic clause := do 190 | atom' ← whnf $ literal.formula $ get_lit c 0, 191 | return $ 192 | if literal.is_neg (get_lit c 0) then 193 | { c with type := atom'.imp (binding_body c.type) } 194 | else 195 | { c with type := `(not %%atom').imp (c.type.binding_body) } 196 | 197 | end clause 198 | 199 | meta def unify_lit (l1 l2 : clause.literal) : tactic unit := 200 | if clause.literal.is_pos l1 = clause.literal.is_pos l2 then 201 | unify (clause.literal.formula l1) (clause.literal.formula l2) transparency.none 202 | else 203 | fail "cannot unify literals" 204 | 205 | -- FIXME: this is most definitely broken with meta-variables that were already in the goal 206 | meta def sort_and_constify_metas : list expr → tactic (list expr) 207 | | exprs_with_metas := do 208 | inst_exprs ← mapm instantiate_mvars exprs_with_metas, 209 | metas ← return $ inst_exprs >>= get_metas, 210 | match list.filter (λm, ¬has_meta_var (get_meta_type m)) metas with 211 | | [] := 212 | if metas.empty then 213 | return [] 214 | else do 215 | metas.mmap' (λm, do trace (expr.to_string m), t ← infer_type m, trace (expr.to_string t)), 216 | fail "could not sort metas" 217 | | ((mvar n pp_n t) :: _) := do 218 | c ← infer_type (mvar n pp_n t) >>= mk_local_def `x, 219 | unify c (mvar n pp_n t), 220 | rest ← sort_and_constify_metas metas, 221 | c ← instantiate_mvars c, 222 | return ([c] ++ rest) 223 | | _ := failed 224 | end 225 | 226 | namespace clause 227 | 228 | meta def meta_closure (metas : list expr) (qf : clause) : tactic clause := do 229 | bs ← sort_and_constify_metas metas, 230 | qf' ← clause.inst_mvars qf, 231 | clause.inst_mvars $ clause.close_constn qf' bs 232 | 233 | private meta def distinct' (local_false : expr) : list expr → expr → clause 234 | | [] proof := ⟨ 0, 0, proof, local_false, local_false ⟩ 235 | | (h::hs) proof := 236 | let (dups, rest) := partition (λh' : expr, h.local_type = h'.local_type) hs, 237 | proof_wo_dups := foldl (λproof (h' : expr), 238 | instantiate_var (abstract_local proof h'.local_uniq_name) h) 239 | proof dups in 240 | (distinct' rest proof_wo_dups).close_const h 241 | 242 | meta def distinct (c : clause) : tactic clause := do 243 | (qf, vs) ← c.open_constn c.num_quants, 244 | (fls, hs) ← qf.open_constn qf.num_lits, 245 | return $ (distinct' c.local_false hs fls.proof).close_constn vs 246 | 247 | end clause 248 | 249 | end super 250 | -------------------------------------------------------------------------------- /src/super/clause_ops.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause 7 | open monad tactic expr 8 | 9 | namespace super 10 | 11 | local attribute [instance] has_monad_lift_to_has_coe 12 | 13 | meta def on_left_at {m} [monad m] (c : clause) (i : ℕ) 14 | [has_monad_lift_t tactic m] 15 | -- f gets a type and returns a list of proofs of that type 16 | (f : expr → m (list (list expr × expr))) : m (list clause) := do 17 | op ← c.open_constn (c.num_quants + i), 18 | @guard tactic _ (op.1.get_lit 0).is_neg _, 19 | new_hyps ← f (op.1.get_lit 0).formula, 20 | return $ new_hyps.map (λnew_hyp, 21 | (op.1.inst new_hyp.2).close_constn (op.2 ++ new_hyp.1)) 22 | 23 | meta def on_left_at_dn {m} [monad m] [alternative m] (c : clause) (i : ℕ) 24 | [has_monad_lift_t tactic m] 25 | -- f gets a hypothesis of ¬type and returns a list of proofs of false 26 | (f : expr → m (list (list expr × expr))) : m (list clause) := do 27 | qf ← c.open_constn c.num_quants, 28 | op ← qf.1.open_constn c.num_lits, 29 | lci ← (op.2.nth i).to_monad, 30 | @guard tactic _ (qf.1.get_lit i).is_neg _, 31 | h ← mk_local_def `h $ imp (qf.1.get_lit i).formula c.local_false, 32 | new_hyps ← f h, 33 | return $ new_hyps.map $ λnew_hyp, 34 | (((clause.mk 0 0 new_hyp.2 c.local_false c.local_false).close_const h).inst 35 | (op.1.close_const lci).proof).close_constn 36 | (qf.2 ++ op.2.remove_nth i ++ new_hyp.1) 37 | 38 | meta def on_right_at {m} [monad m] (c : clause) (i : ℕ) 39 | [has_monad_lift_t tactic m] 40 | -- f gets a hypothesis and returns a list of proofs of false 41 | (f : expr → m (list (list expr × expr))) : m (list clause) := do 42 | op ← c.open_constn (c.num_quants + i), 43 | @guard tactic _ ((op.1.get_lit 0).is_pos) _, 44 | h ← mk_local_def `h (op.1.get_lit 0).formula, 45 | new_hyps ← f h, 46 | return $ new_hyps.map (λnew_hyp, 47 | (op.1.inst (lambdas [h] new_hyp.2)).close_constn (op.2 ++ new_hyp.1)) 48 | 49 | meta def on_right_at' {m} [monad m] (c : clause) (i : ℕ) 50 | [has_monad_lift_t tactic m] 51 | -- f gets a hypothesis and returns a list of proofs 52 | (f : expr → m (list (list expr × expr))) : m (list clause) := do 53 | op ← c.open_constn (c.num_quants + i), 54 | @guard tactic _ ((op.1.get_lit 0).is_pos) _, 55 | h ← mk_local_def `h (op.1.get_lit 0).formula, 56 | new_hyps ← f h, 57 | new_hyps.mmap (λnew_hyp, do 58 | type ← infer_type new_hyp.2, 59 | nh ← mk_local_def `nh $ imp type c.local_false, 60 | return $ (op.1.inst (lambdas [h] (app nh new_hyp.2))).close_constn (op.2 ++ new_hyp.1 ++ [nh])) 61 | 62 | meta def on_first_right (c : clause) (f : expr → tactic (list (list expr × expr))) : tactic (list clause) := 63 | first $ do i ← list.range c.num_lits, [on_right_at c i f] 64 | 65 | meta def on_first_right' (c : clause) (f : expr → tactic (list (list expr × expr))) : tactic (list clause) := 66 | first $ do i ← list.range c.num_lits, [on_right_at' c i f] 67 | 68 | meta def on_first_left (c : clause) (f : expr → tactic (list (list expr × expr))) : tactic (list clause) := 69 | first $ do i ← list.range c.num_lits, [on_left_at c i f] 70 | 71 | meta def on_first_left_dn (c : clause) (f : expr → tactic (list (list expr × expr))) : tactic (list clause) := 72 | first $ do i ← list.range c.num_lits, [on_left_at_dn c i f] 73 | 74 | end super 75 | -------------------------------------------------------------------------------- /src/super/clausifier.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause .clause_ops 7 | import .prover_state .misc_preprocessing 8 | open expr list tactic monad decidable 9 | 10 | universe u 11 | 12 | namespace super 13 | 14 | meta def try_option {a : Type u} (tac : tactic a) : tactic (option a) := 15 | some <$> tac <|> return none 16 | 17 | private meta def normalize : expr → tactic expr | e := do 18 | e' ← whnf e reducible, 19 | args' ← e'.get_app_args.mmap normalize, 20 | return $ app_of_list e'.get_app_fn args' 21 | 22 | meta def inf_normalize_l (c : clause) : tactic (list clause) := 23 | on_first_left c $ λtype, do 24 | type' ← normalize type, 25 | guard $ type' ≠ type, 26 | h ← mk_local_def `h type', 27 | return [([h], h)] 28 | 29 | meta def inf_normalize_r (c : clause) : tactic (list clause) := 30 | on_first_right c $ λha, do 31 | a' ← normalize ha.local_type, 32 | guard $ a' ≠ ha.local_type, 33 | hna ← mk_local_def `hna (imp a' c.local_false), 34 | return [([hna], app hna ha)] 35 | 36 | meta def inf_false_l (c : clause) : tactic (list clause) := 37 | first $ do i ← list.range c.num_lits, 38 | if c.get_lit i = clause.literal.left `(false) 39 | then [return []] 40 | else [] 41 | 42 | meta def inf_false_r (c : clause) : tactic (list clause) := 43 | on_first_right c $ λhf, 44 | if hf.local_type = c.local_false 45 | then return [([], hf)] 46 | else match hf.local_type with 47 | | const ``false [] := do 48 | pr ← mk_app ``false.rec [c.local_false, hf], 49 | return [([], pr)] 50 | | _ := failed 51 | end 52 | 53 | meta def inf_true_l (c : clause) : tactic (list clause) := 54 | on_first_left c $ λt, 55 | match t with 56 | | (const ``true []) := return [([], const ``true.intro [])] 57 | | _ := failed 58 | end 59 | 60 | meta def inf_true_r (c : clause) : tactic (list clause) := 61 | first $ do i ← list.range c.num_lits, 62 | if c.get_lit i = clause.literal.right (const ``true []) 63 | then [return []] 64 | else [] 65 | 66 | meta def inf_not_l (c : clause) : tactic (list clause) := 67 | on_first_left c $ λtype, 68 | match type with 69 | | app (const ``not []) a := do 70 | hna ← mk_local_def `h (a.imp `(false)), 71 | return [([hna], hna)] 72 | | _ := failed 73 | end 74 | 75 | meta def inf_not_r (c : clause) : tactic (list clause) := 76 | on_first_right c $ λhna, 77 | match hna.local_type with 78 | | app (const ``not []) a := do 79 | hnna ← mk_local_def `h ((a.imp `(false)).imp c.local_false), 80 | return [([hnna], app hnna hna)] 81 | | _ := failed 82 | end 83 | 84 | meta def inf_and_l (c : clause) : tactic (list clause) := 85 | on_first_left c $ λab, 86 | match ab with 87 | | (app (app (const ``and []) a) b) := do 88 | ha ← mk_local_def `l a, 89 | hb ← mk_local_def `r b, 90 | pab ← mk_mapp ``and.intro [some a, some b, some ha, some hb], 91 | return [([ha, hb], pab)] 92 | | _ := failed 93 | end 94 | 95 | meta def inf_and_r (c : clause) : tactic (list clause) := 96 | on_first_right' c $ λhyp, do 97 | pa ← mk_mapp ``and.left [none, none, some hyp], 98 | pb ← mk_mapp ``and.right [none, none, some hyp], 99 | return [([], pa), ([], pb)] 100 | 101 | meta def inf_iff_l (c : clause) : tactic (list clause) := 102 | on_first_left c $ λab, 103 | match ab with 104 | | (app (app (const ``iff []) a) b) := do 105 | hab ← mk_local_def `l (imp a b), 106 | hba ← mk_local_def `r (imp b a), 107 | pab ← mk_mapp ``iff.intro [some a, some b, some hab, some hba], 108 | return [([hab, hba], pab)] 109 | | _ := failed 110 | end 111 | 112 | meta def inf_iff_r (c : clause) : tactic (list clause) := 113 | on_first_right' c $ λhyp, do 114 | pa ← mk_mapp ``iff.mp [none, none, some hyp], 115 | pb ← mk_mapp ``iff.mpr [none, none, some hyp], 116 | return [([], pa), ([], pb)] 117 | 118 | meta def inf_or_r (c : clause) : tactic (list clause) := 119 | on_first_right c $ λhab, 120 | match hab.local_type with 121 | | (app (app (const ``or []) a) b) := do 122 | hna ← mk_local_def `l (imp a c.local_false), 123 | hnb ← mk_local_def `r (imp b c.local_false), 124 | proof ← mk_app ``or.elim [a, b, c.local_false, hab, hna, hnb], 125 | return [([hna, hnb], proof)] 126 | | _ := failed 127 | end 128 | 129 | meta def inf_or_l (c : clause) : tactic (list clause) := 130 | on_first_left c $ λab, 131 | match ab with 132 | | (app (app (const ``or []) a) b) := do 133 | ha ← mk_local_def `l a, 134 | hb ← mk_local_def `l b, 135 | pa ← mk_mapp ``or.inl [some a, some b, some ha], 136 | pb ← mk_mapp ``or.inr [some a, some b, some hb], 137 | return [([ha], pa), ([hb], pb)] 138 | | _ := failed 139 | end 140 | 141 | meta def inf_all_r (c : clause) : tactic (list clause) := 142 | on_first_right' c $ λhallb, 143 | match hallb.local_type with 144 | | (pi n bi a b) := do 145 | ha ← mk_local_def `x a, 146 | return [([ha], app hallb ha)] 147 | | _ := failed 148 | end 149 | 150 | lemma imp_l {F a b} [decidable a] : ((a → b) → F) → ((a → F) → F) := 151 | λhabf haf, decidable.by_cases 152 | (assume ha : a, haf ha) 153 | (assume hna : ¬a, habf (assume ha, absurd ha hna)) 154 | 155 | lemma imp_l' {F a b} [decidable F] : ((a → b) → F) → ((a → F) → F) := 156 | λhabf haf, decidable.by_cases 157 | (assume hf : F, hf) 158 | (assume hnf : ¬F, habf (assume ha, absurd (haf ha) hnf)) 159 | 160 | lemma imp_l_c {F : Prop} {a b} : ((a → b) → F) → ((a → F) → F) := 161 | λhabf haf, classical.by_cases 162 | (assume hf : F, hf) 163 | (assume hnf : ¬F, habf (assume ha, absurd (haf ha) hnf)) 164 | 165 | meta def inf_imp_l (c : clause) : tactic (list clause) := 166 | on_first_left_dn c $ λhnab, 167 | match hnab.local_type with 168 | | (pi _ _ (pi _ _ a b) _) := 169 | if b.has_var then failed else do 170 | hna ← mk_local_def `na (imp a c.local_false), 171 | pf ← first (do r ← [``super.imp_l, ``super.imp_l', ``super.imp_l_c], 172 | [mk_app r [hnab, hna]]), 173 | hb ← mk_local_def `b b, 174 | return [([hna], pf), ([hb], app hnab (lam `a binder_info.default a hb))] 175 | | _ := failed 176 | end 177 | 178 | meta def inf_ex_l (c : clause) : tactic (list clause) := 179 | on_first_left c $ λexp, 180 | match exp with 181 | | (app (app (const ``Exists [u]) dom) pred) := do 182 | hx ← mk_local_def `x dom, 183 | predx ← whnf $ app pred hx, 184 | hpx ← mk_local_def `hpx predx, 185 | return [([hx,hpx], app_of_list (const ``exists.intro [u]) 186 | [dom, pred, hx, hpx])] 187 | | _ := failed 188 | end 189 | 190 | lemma demorgan' {F a} {b : a → Prop} : ((∀x, b x) → F) → (((∃x, b x → F) → F) → F) := 191 | assume hab hnenb, 192 | classical.by_cases 193 | (assume h : ∃x, ¬b x, begin cases h with x, apply hnenb, existsi x, intros, contradiction end) 194 | (assume h : ¬∃x, ¬b x, hab (assume x, 195 | classical.by_cases 196 | (assume bx : b x, bx) 197 | (assume nbx : ¬b x, have hf : false, { apply h, existsi x, assumption }, by contradiction))) 198 | 199 | meta def inf_all_l (c : clause) : tactic (list clause) := 200 | on_first_left_dn c $ λhnallb, 201 | match hnallb.local_type with 202 | | pi _ _ (pi n bi a b) _ := do 203 | enb ← mk_mapp ``Exists [none, some $ lam n binder_info.default a (imp b c.local_false)], 204 | hnenb ← mk_local_def `h (imp enb c.local_false), 205 | pr ← mk_app ``super.demorgan' [hnallb, hnenb], 206 | return [([hnenb], pr)] 207 | | _ := failed 208 | end 209 | 210 | meta def inf_ex_r (c : clause) : tactic (list clause) := do 211 | (qf, ctx) ← c.open_constn c.num_quants, 212 | skolemized ← on_first_right' qf $ λhexp, 213 | match hexp.local_type with 214 | | (app (app (const ``Exists [_]) d) p) := do 215 | sk_sym_name_pp ← get_unused_name `sk (some 1), 216 | inh_lc ← mk_local' `w binder_info.implicit d, 217 | sk_sym ← mk_local_def sk_sym_name_pp (pis (ctx ++ [inh_lc]) d), 218 | sk_p ← whnf_no_delta $ app p (app_of_list sk_sym (ctx ++ [inh_lc])), 219 | sk_ax ← mk_mapp ``Exists [some (local_type sk_sym), 220 | some (lambdas [sk_sym] (pis (ctx ++ [inh_lc]) (imp hexp.local_type sk_p)))], 221 | sk_ax_name ← get_unused_name `sk_axiom (some 1), assert sk_ax_name sk_ax, 222 | nonempt_of_inh ← mk_mapp ``nonempty.intro [some d, some inh_lc], 223 | eps ← mk_mapp ``classical.epsilon [some d, some nonempt_of_inh, some p], 224 | existsi (lambdas (ctx ++ [inh_lc]) eps), 225 | eps_spec ← mk_mapp ``classical.epsilon_spec [some d, some p], 226 | exact (lambdas (ctx ++ [inh_lc]) eps_spec), 227 | sk_ax_local ← get_local sk_ax_name, cases sk_ax_local [sk_sym_name_pp, sk_ax_name], 228 | sk_ax' ← get_local sk_ax_name, 229 | return [([inh_lc], app_of_list sk_ax' (ctx ++ [inh_lc, hexp]))] 230 | | _ := failed 231 | end, 232 | return $ skolemized.map (λs, s.close_constn ctx) 233 | 234 | meta def first_some {a : Type} : list (tactic (option a)) → tactic (option a) 235 | | [] := return none 236 | | (x::xs) := do xres ← x, match xres with some y := return (some y) | none := first_some xs end 237 | 238 | private meta def get_clauses_core' (rules : list (clause → tactic (list clause))) 239 | : list clause → tactic (list clause) | cs := 240 | list.join <$> do 241 | cs.mmap $ λc, do first $ 242 | rules.map (λr, r c >>= get_clauses_core') ++ [return [c]] 243 | 244 | meta def get_clauses_core (rules : list (clause → tactic (list clause))) (initial : list clause) 245 | : tactic (list clause) := do 246 | clauses ← get_clauses_core' rules initial, 247 | filter (λc, bnot <$> is_taut c) $ list.nub_on clause.type clauses 248 | 249 | meta def clausification_rules_intuit : list (clause → tactic (list clause)) := 250 | [ inf_false_l, inf_false_r, inf_true_l, inf_true_r, 251 | inf_not_l, inf_not_r, 252 | inf_and_l, inf_and_r, 253 | inf_iff_l, inf_iff_r, 254 | inf_or_l, inf_or_r, 255 | inf_ex_l, 256 | inf_normalize_l, inf_normalize_r ] 257 | 258 | meta def clausification_rules_classical : list (clause → tactic (list clause)) := 259 | [ inf_false_l, inf_false_r, inf_true_l, inf_true_r, 260 | inf_not_l, inf_not_r, 261 | inf_and_l, inf_and_r, 262 | inf_iff_l, inf_iff_r, 263 | inf_or_l, inf_or_r, 264 | inf_imp_l, inf_all_r, 265 | inf_ex_l, 266 | inf_all_l, inf_ex_r, 267 | inf_normalize_l, inf_normalize_r ] 268 | 269 | meta def get_clauses_classical : list clause → tactic (list clause) := 270 | get_clauses_core clausification_rules_classical 271 | meta def get_clauses_intuit : list clause → tactic (list clause) := 272 | get_clauses_core clausification_rules_intuit 273 | 274 | meta def as_refutation : tactic unit := do 275 | repeat (do intro1, skip), 276 | tgt ← target, 277 | if tgt.is_constant || tgt.is_local_constant then skip else do 278 | local_false_name ← get_unused_name `F none, tgt_type ← infer_type tgt, 279 | definev local_false_name tgt_type tgt, local_false ← get_local local_false_name, 280 | target_name ← get_unused_name `target none, 281 | assertv target_name (imp tgt local_false) (lam `hf binder_info.default tgt $ mk_var 0), 282 | change local_false 283 | 284 | meta def clauses_of_context : tactic (list clause) := do 285 | local_false ← target, 286 | l ← local_context, 287 | l.mmap (clause.of_proof local_false) 288 | 289 | meta def clausify_pre := preprocessing_rule $ assume new, list.join <$> new.mmap (λ dc, do 290 | cs ← get_clauses_classical [dc.c], 291 | if cs.length ≤ 1 then 292 | return (cs.map $ λ c, { dc with c := c }) 293 | else 294 | cs.mmap (λc, mk_derived c dc.sc)) 295 | 296 | -- @[super.inf] 297 | meta def clausification_inf : inf_decl := inf_decl.mk 0 $ 298 | λ given, list.foldr (<|>) (return ()) $ 299 | do r ← clausification_rules_classical, 300 | [do cs ← r given.c, 301 | cs' ← get_clauses_classical cs, 302 | cs'.mmap' (λc, mk_derived c given.sc.sched_now >>= add_inferred), 303 | remove_redundant given.id []] 304 | 305 | 306 | end super 307 | -------------------------------------------------------------------------------- /src/super/datatypes.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause_ops .prover_state 7 | open expr tactic monad 8 | 9 | namespace super 10 | 11 | meta def has_diff_constr_eq_l (c : clause) : tactic bool := do 12 | env ← get_env, 13 | return $ list.bor $ do 14 | l ← c.get_lits, 15 | guard l.is_neg, 16 | return $ match is_eq l.formula with 17 | | some (lhs, rhs) := 18 | if env.is_constructor_app lhs ∧ env.is_constructor_app rhs ∧ 19 | lhs.get_app_fn.const_name ≠ rhs.get_app_fn.const_name then 20 | tt 21 | else 22 | ff 23 | | _ := ff 24 | end 25 | 26 | meta def diff_constr_eq_l_pre := preprocessing_rule $ 27 | filter (λc, bnot <$> has_diff_constr_eq_l c.c) 28 | 29 | meta def try_no_confusion_eq_r (c : clause) (i : ℕ) : tactic (list clause) := 30 | on_right_at' c i $ λhyp, 31 | match is_eq hyp.local_type with 32 | | some (lhs, rhs) := do 33 | env ← get_env, 34 | lhs ← whnf lhs, rhs ← whnf rhs, 35 | guard $ env.is_constructor_app lhs ∧ env.is_constructor_app rhs, 36 | pr ← mk_app (lhs.get_app_fn.const_name.get_prefix <.> "no_confusion") [`(false), lhs, rhs, hyp], 37 | -- FIXME: change to local false ^^ 38 | ty ← infer_type pr, ty ← whnf ty, 39 | pr ← to_expr ``(@eq.mpr _ %%ty rfl %%pr), -- FIXME 40 | return [([], pr)] 41 | | _ := failed 42 | end 43 | 44 | @[super.inf] 45 | meta def datatype_infs : inf_decl := inf_decl.mk 10 $ assume given, do 46 | sequence' (do i ← list.range given.c.num_lits, [inf_if_successful 0 given $ try_no_confusion_eq_r given.c i]) 47 | 48 | end super 49 | -------------------------------------------------------------------------------- /src/super/default.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .prover 7 | -------------------------------------------------------------------------------- /src/super/defs.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause_ops .prover_state 7 | open tactic expr monad 8 | 9 | namespace super 10 | 11 | meta def try_unfold_one_def (type : expr) : tactic expr := 12 | dunfold_head type transparency.all 13 | 14 | meta def try_unfold_def_left (c : clause) (i : ℕ) : tactic (list clause) := 15 | on_left_at c i $ λt, do 16 | t' ← try_unfold_one_def t, 17 | ht' ← mk_local_def `h t', 18 | return [([ht'], ht')] 19 | 20 | meta def try_unfold_def_right (c : clause) (i : ℕ) : tactic (list clause) := 21 | on_right_at c i $ λh, do 22 | t' ← try_unfold_one_def h.local_type, 23 | hnt' ← mk_local_def `h (imp t' c.local_false), 24 | return [([hnt'], app hnt' h)] 25 | 26 | -- @[super.inf] 27 | meta def unfold_def_inf : inf_decl := inf_decl.mk 40 $ assume given, sequence' $ do 28 | r ← [try_unfold_def_right, try_unfold_def_left], 29 | -- NOTE: we cannot restrict to selected literals here 30 | -- as this might prevent factoring, e.g. _n>0_ ∨ is_pos(0) 31 | i ← list.range given.c.num_lits, 32 | [inf_if_successful 3 given (r given.c i)] 33 | 34 | end super 35 | -------------------------------------------------------------------------------- /src/super/demod.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .superposition 7 | open tactic monad expr 8 | 9 | namespace super 10 | 11 | meta def is_demodulator (c : clause) : bool := 12 | match c.get_lits with 13 | | [clause.literal.right eqn] := eqn.is_eq.is_some 14 | | _ := ff 15 | end 16 | 17 | variable gt : expr → expr → bool 18 | 19 | meta def demod' (cs1 : list derived_clause) : clause → list derived_clause → tactic (list derived_clause × clause) | c2 used_demods := 20 | (first $ do 21 | i ← list.range c2.num_lits, 22 | pos ← rwr_positions c2 i, 23 | c1 ← cs1, 24 | (ltr, congr_ax) ← [(tt, ``super.sup_ltr), (ff, ``super.sup_rtl)], 25 | [do c2' ← try_sup gt c1.c c2 0 i pos ltr tt congr_ax, 26 | demod' c2' (c1 :: used_demods)] 27 | ) <|> return (used_demods, c2) 28 | 29 | meta def demod (cs1 : list derived_clause) (c2 : clause) : tactic (list derived_clause × clause) := do 30 | c2qf ← c2.open_constn c2.num_quants, 31 | (used_demods, c2qf') ← demod' gt cs1 c2qf.1 [], 32 | if used_demods.empty then return ([], c2) else 33 | return (used_demods, c2qf'.close_constn c2qf.2) 34 | 35 | meta def demod_fwd_inf : inference := 36 | assume given, do 37 | active ← get_active, 38 | demods ← return (do ac ← active.values, guard $ is_demodulator ac.c, guard $ ac.id ≠ given.id, [ac]), 39 | if demods.empty then skip else do 40 | (used_demods, given') ← demod gt demods given.c, 41 | if used_demods.empty then skip else do 42 | remove_redundant given.id used_demods, 43 | mk_derived given' given.sc.sched_now >>= add_inferred 44 | 45 | meta def demod_back1 (given active : derived_clause) : prover unit := do 46 | (used_demods, c') ← demod gt [given] active.c, 47 | if used_demods.empty then skip else do 48 | remove_redundant active.id used_demods, 49 | mk_derived c' active.sc.sched_now >>= add_inferred 50 | 51 | meta def demod_back_inf : inference := 52 | assume given, if ¬is_demodulator given.c then skip else 53 | do active ← get_active, sequence' $ do 54 | ac ← active.values, 55 | guard $ ac.id ≠ given.id, 56 | [demod_back1 gt given ac] 57 | 58 | @[super.inf] 59 | meta def demod_inf : inf_decl := { 60 | prio := 10, 61 | inf := assume given, do 62 | gt ← get_term_order, 63 | demod_fwd_inf gt given, 64 | demod_back_inf gt given, 65 | skip 66 | } 67 | 68 | end super 69 | -------------------------------------------------------------------------------- /src/super/equality.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause .prover_state .utils 7 | open tactic monad expr list 8 | 9 | namespace super 10 | 11 | meta def try_unify_eq_l (c : clause) (i : nat) : tactic clause := do 12 | guard $ clause.literal.is_neg (clause.get_lit c i), 13 | qf ← clause.open_metan c c.num_quants, 14 | match is_eq (qf.1.get_lit i).formula with 15 | | none := failed 16 | | some (lhs, rhs) := do 17 | unify lhs rhs, 18 | ty ← infer_type lhs, 19 | univ ← infer_univ ty, 20 | refl ← return $ app_of_list (const ``eq.refl [univ]) [ty, lhs], 21 | opened ← clause.open_constn qf.1 i, 22 | clause.meta_closure qf.2 $ clause.close_constn (opened.1.inst refl) opened.2 23 | end 24 | 25 | @[super.inf] 26 | meta def unify_eq_inf : inf_decl := inf_decl.mk 40 $ assume given, sequence' $ do 27 | i ← given.selected, 28 | [inf_if_successful 0 given (do u ← try_unify_eq_l given.c i, return [u])] 29 | 30 | meta def has_refl_r (c : clause) : bool := 31 | list.bor $ do 32 | literal ← c.get_lits, 33 | guard literal.is_pos, 34 | match is_eq literal.formula with 35 | | some (lhs, rhs) := [decidable.to_bool (lhs = rhs)] 36 | | none := [] 37 | end 38 | 39 | meta def refl_r_pre : prover unit := 40 | preprocessing_rule $ assume new, return $ filter (λc, ¬has_refl_r c.c) new 41 | 42 | end super 43 | -------------------------------------------------------------------------------- /src/super/factoring.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause .prover_state .subsumption 7 | open tactic expr monad 8 | 9 | namespace super 10 | 11 | variable gt : expr → expr → bool 12 | 13 | meta def inst_lit (c : clause) (i : nat) (e : expr) : tactic clause := do 14 | opened ← clause.open_constn c i, 15 | return $ clause.close_constn (clause.inst opened.1 e) opened.2 16 | 17 | private meta def try_factor' (c : clause) (i j : nat) : tactic clause := do 18 | -- instantiate universal quantifiers using meta-variables 19 | (qf, mvars) ← c.open_metan c.num_quants, 20 | -- unify the two literals 21 | unify_lit (qf.get_lit i) (qf.get_lit j), 22 | -- check maximality condition 23 | qfi ← qf.inst_mvars, guard $ clause.is_maximal gt qfi i, 24 | -- construct proof 25 | (at_j, cs) ← qf.open_constn j, hyp_i ← cs.nth i, 26 | let qf' := (at_j.inst hyp_i).close_constn cs, 27 | -- instantiate meta-variables and replace remaining meta-variables by quantifiers 28 | clause.meta_closure mvars qf' 29 | 30 | meta def try_factor (c : clause) (i j : nat) : tactic clause := 31 | if i > j then try_factor' gt c j i else try_factor' gt c i j 32 | 33 | meta def try_infer_factor (c : derived_clause) (i j : nat) : prover unit := do 34 | f ← try_factor gt c.c i j, 35 | ss ← does_subsume f c.c, 36 | if ss then do 37 | f ← mk_derived f c.sc.sched_now, 38 | add_inferred f, 39 | remove_redundant c.id [f] 40 | else do 41 | inf_score 1 [c.sc] >>= mk_derived f >>= add_inferred 42 | 43 | @[super.inf] 44 | meta def factor_inf : inf_decl := inf_decl.mk 40 $ 45 | assume given, do gt ← get_term_order, sequence' $ do 46 | i ← given.selected, 47 | j ← list.range given.c.num_lits, 48 | return $ try_infer_factor gt given i j <|> return () 49 | 50 | meta def factor_dup_lits_pre := preprocessing_rule $ assume new, do 51 | new.mmap $ λdc, do 52 | dist ← dc.c.distinct, 53 | return { dc with c := dist } 54 | 55 | end super 56 | -------------------------------------------------------------------------------- /src/super/inhabited.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause_ops .prover_state 7 | open expr tactic monad 8 | 9 | namespace super 10 | 11 | meta def try_assumption_lookup_left (c : clause) : tactic (list clause) := 12 | on_first_left c $ λtype, do 13 | ass ← find_assumption type, 14 | return [([], ass)] 15 | 16 | meta def try_nonempty_lookup_left (c : clause) : tactic (list clause) := 17 | on_first_left_dn c $ λhnx, 18 | match is_local_not c.local_false hnx.local_type with 19 | | some type := do 20 | univ ← infer_univ type, 21 | lf_univ ← infer_univ c.local_false, 22 | guard $ lf_univ = level.zero, 23 | inst ← mk_instance (app (const ``nonempty [univ]) type), 24 | instt ← infer_type inst, 25 | return [([], app_of_list (const ``nonempty.elim [univ]) 26 | [type, c.local_false, inst, hnx])] 27 | | _ := failed 28 | end 29 | 30 | meta def try_nonempty_left (c : clause) : tactic (list clause) := 31 | on_first_left c $ λprop, 32 | match prop with 33 | | (app (const ``nonempty [u]) type) := do 34 | x ← mk_local_def `x type, 35 | return [([x], app_of_list (const ``nonempty.intro [u]) [type, x])] 36 | | _ := failed 37 | end 38 | 39 | meta def try_nonempty_right (c : clause) : tactic (list clause) := 40 | on_first_right c $ λhnonempty, 41 | match hnonempty.local_type with 42 | | (app (const ``nonempty [u]) type) := do 43 | lf_univ ← infer_univ c.local_false, 44 | guard $ lf_univ = level.zero, 45 | hnx ← mk_local_def `nx (imp type c.local_false), 46 | return [([hnx], app_of_list (const ``nonempty.elim [u]) 47 | [type, c.local_false, hnonempty, hnx])] 48 | | _ := failed 49 | end 50 | 51 | meta def try_inhabited_left (c : clause) : tactic (list clause) := 52 | on_first_left c $ λprop, 53 | match prop with 54 | | (app (const ``inhabited [u]) type) := do 55 | x ← mk_local_def `x type, 56 | return [([x], app_of_list (const ``inhabited.mk [u]) [type, x])] 57 | | _ := failed 58 | end 59 | 60 | meta def try_inhabited_right (c : clause) : tactic (list clause) := 61 | on_first_right' c $ λhinh, 62 | match hinh.local_type with 63 | | (app (const ``inhabited [u]) type) := 64 | return [([], app_of_list (const ``inhabited.default [u]) [type, hinh])] 65 | | _ := failed 66 | end 67 | 68 | @[super.inf] 69 | meta def inhabited_infs : inf_decl := inf_decl.mk 10 $ assume given, do 70 | [try_assumption_lookup_left, 71 | try_nonempty_lookup_left, 72 | try_nonempty_left, try_nonempty_right, 73 | try_inhabited_left, try_inhabited_right].mmap' $ λr, 74 | simp_if_successful given (r given.c) 75 | 76 | end super 77 | -------------------------------------------------------------------------------- /src/super/lpo.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | 7 | -- Polytime version of lexicographic path order as described in: 8 | -- Things to know when implementing LPO, Bernd Löchner, ESFOR 2004 9 | import .utils 10 | open expr decidable monad native 11 | 12 | def lex {T} [decidable_eq T] (gt : T → T → bool) : list T → list T → bool 13 | | (s::ss) (t::ts) := if s = t then lex ss ts else gt s t 14 | | _ _ := ff 15 | 16 | def majo {T} (gt : T → T → bool) (s : T) : list T → bool 17 | | [] := tt 18 | | (t::ts) := gt s t && majo ts 19 | 20 | meta def alpha (lpo : expr → expr → bool) : list expr → expr → bool 21 | | [] _ := ff 22 | | (s::ss) t := (s = t) || lpo s t || alpha ss t 23 | 24 | meta def lex_ma (lpo : expr → expr → bool) (s t : expr) : list expr → list expr → bool 25 | | (si::ss) (ti::ts) := 26 | if si = ti then lex_ma ss ts 27 | else if lpo si ti then majo lpo s ts 28 | else alpha lpo (si::ss) t 29 | | _ _ := ff 30 | 31 | meta def lpo (prec_gt : expr → expr → bool) : expr → expr → bool | s t := 32 | if prec_gt (get_app_fn s) (get_app_fn t) then majo lpo s (get_app_args t) 33 | else if get_app_fn s = get_app_fn t then lex_ma lpo s t (get_app_args s) (get_app_args t) 34 | else alpha lpo (get_app_args s) t 35 | 36 | meta def prec_gt_of_name_list (ns : list name) : expr → expr → bool := 37 | let nis := rb_map.of_list ns.zip_with_index in 38 | λs t, match (nis.find (name_of_funsym s), nis.find (name_of_funsym t)) with 39 | | (some si, some ti) := si > ti 40 | | _ := ff 41 | end 42 | 43 | meta def mk_lpo (ns : list name) : expr → expr → bool := 44 | let prec_gt := prec_gt_of_name_list ns in 45 | lpo (prec_gt_of_name_list ns) 46 | -------------------------------------------------------------------------------- /src/super/misc_preprocessing.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause .prover_state 7 | open expr list monad 8 | 9 | namespace super 10 | 11 | meta def is_taut (c : clause) : tactic bool := do 12 | qf ← c.open_constn c.num_quants, 13 | return $ list.bor $ do 14 | l1 ← qf.1.get_lits, guard l1.is_neg, 15 | l2 ← qf.1.get_lits, guard l2.is_pos, 16 | [decidable.to_bool $ l1.formula = l2.formula] 17 | 18 | meta def tautology_removal_pre : prover unit := 19 | preprocessing_rule $ λnew, filter (λc, bnot <$> is_taut c.c) new 20 | 21 | meta def remove_duplicates : list derived_clause → list derived_clause 22 | | [] := [] 23 | | (c :: cs) := 24 | let (same_type, other_type) := partition (λc' : derived_clause, c'.c.type = c.c.type) cs in 25 | { c with sc := foldl score.min c.sc (same_type.map $ λc, c.sc) } :: remove_duplicates other_type 26 | 27 | meta def remove_duplicates_pre : prover unit := 28 | preprocessing_rule $ λnew, 29 | return $ remove_duplicates new 30 | 31 | meta def clause_normalize_pre : prover unit := 32 | preprocessing_rule $ λnew, new.mmap $ λdc, 33 | do c' ← dc.c.normalize, return { dc with c := c' } 34 | 35 | end super 36 | -------------------------------------------------------------------------------- /src/super/prover.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause .prover_state 7 | import .misc_preprocessing 8 | import .selection 9 | import .trim 10 | 11 | -- default inferences 12 | -- 0 13 | import .clausifier 14 | -- 10 15 | import .demod 16 | import .inhabited 17 | import .datatypes 18 | -- 20 19 | import .subsumption 20 | -- 30 21 | import .splitting 22 | -- 40 23 | import .factoring 24 | import .resolution 25 | import .superposition 26 | import .equality 27 | import .simp 28 | import .defs 29 | 30 | open monad tactic expr 31 | 32 | declare_trace super 33 | 34 | namespace super 35 | 36 | meta def trace_clauses : prover unit := 37 | do state ← get, trace state 38 | 39 | meta def run_prover_loop 40 | (literal_selection : selection_strategy) 41 | (clause_selection : clause_selection_strategy) 42 | (preprocessing_rules : list (prover unit)) 43 | (inference_rules : list inference) 44 | (max_iters : ℕ) 45 | : ℕ → prover (option expr) | i := do 46 | sequence' preprocessing_rules, 47 | new ← take_newly_derived, new.mmap' register_as_passive, 48 | when (is_trace_enabled_for `super) $ new.mmap' $ λn, 49 | tactic.trace { n with c := { n.c with proof := const (mk_simple_name "derived") [] } }, 50 | needs_sat_run ← flip (<$>) get (λst, st.needs_sat_run), 51 | if needs_sat_run then do 52 | res ← do_sat_run, 53 | match res with 54 | | some proof := return (some proof) 55 | | none := do 56 | model ← flip (<$>) get (λst, st.current_model), 57 | when (is_trace_enabled_for `super) (do 58 | pp_model ← pp (model.to_list.map (λlit, if lit.2 = tt then lit.1 else `(not %%lit.1))), 59 | trace $ to_fmt "sat model: " ++ pp_model), 60 | run_prover_loop i 61 | end 62 | else do 63 | passive ← get_passive, 64 | if passive.size = 0 then return none else do 65 | given_name ← clause_selection i, 66 | given ← option.to_monad (passive.find given_name), 67 | -- trace_clauses, 68 | remove_passive given_name, 69 | given ← literal_selection given, 70 | when (is_trace_enabled_for `super) (do 71 | fmt ← pp given, trace (to_fmt "given: " ++ fmt)), 72 | add_active given, 73 | seq_inferences inference_rules given, 74 | when (i ≥ max_iters) $ fail "reached maximum number of iterations in saturation loop", 75 | run_prover_loop (i+1) 76 | 77 | meta def default_preprocessing : list (prover unit) := 78 | [ 79 | clausify_pre, 80 | clause_normalize_pre, 81 | factor_dup_lits_pre, 82 | remove_duplicates_pre, 83 | refl_r_pre, 84 | diff_constr_eq_l_pre, 85 | tautology_removal_pre, 86 | subsumption_interreduction_pre, 87 | forward_subsumption_pre, 88 | return () 89 | ] 90 | 91 | end super 92 | 93 | open super 94 | 95 | meta structure super.opts := 96 | (max_iters : ℕ := 200) 97 | (timeout : ℕ := 20000) 98 | 99 | meta def super (sos_lemmas : list expr) (os : super.opts) : tactic unit := do 100 | try_for os.timeout $ with_trim $ do 101 | as_refutation, local_false ← target, 102 | clauses ← clauses_of_context, 103 | sos_clauses ← sos_lemmas.mmap (clause.of_proof local_false), 104 | initial_state ← prover_state.initial local_false (clauses ++ sos_clauses), 105 | inf_names ← attribute.get_instances `super.inf, 106 | infs ← inf_names.mmap $ λn, eval_expr inf_decl (const n []), 107 | infs ← return $ list.map inf_decl.inf $ list.sort_on inf_decl.prio infs, 108 | res ← (run_prover_loop selection21 (age_weight_clause_selection 3 4) 109 | default_preprocessing infs 110 | os.max_iters 0).run initial_state, 111 | match res with 112 | | (some empty_clause, st) := apply empty_clause >> skip 113 | | (none, saturation) := do sat_fmt ← pp saturation, 114 | fail $ to_fmt "saturation:" ++ format.line ++ sat_fmt 115 | end 116 | 117 | namespace tactic.interactive 118 | open lean.parser 119 | open interactive 120 | open interactive.types 121 | 122 | meta def with_lemmas (ls : parse $ many ident) : tactic unit := 123 | ls.mmap' $ λ l, do 124 | eqn_ls ← get_eqn_lemmas_for tt l <|> return [], 125 | (l :: eqn_ls).mmap' $ λ l, do 126 | p ← mk_const l, 127 | t ← infer_type p, 128 | n ← get_unused_name p.get_app_fn.const_name none, 129 | tactic.assertv n t p 130 | 131 | meta def super (extra_clause_names : parse $ many ident) 132 | (extra_lemma_names : parse with_ident_list) 133 | (opts : super.opts := {}) : tactic unit := do 134 | with_lemmas extra_clause_names, 135 | extra_lemmas ← extra_lemma_names.mmap (λ c, do 136 | eqn_ls ← get_eqn_lemmas_for tt c, 137 | (c::eqn_ls).mmap mk_const), 138 | _root_.super extra_lemmas.join opts 139 | 140 | end tactic.interactive 141 | -------------------------------------------------------------------------------- /src/super/prover_state.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause .lpo .cdcl_solver 7 | open tactic functor monad expr native 8 | 9 | namespace super 10 | 11 | structure score := 12 | (priority : ℕ) 13 | (in_sos : bool) 14 | (cost : ℕ) 15 | (age : ℕ) 16 | 17 | namespace score 18 | def prio.immediate : ℕ := 0 19 | def prio.default : ℕ := 1 20 | def prio.never : ℕ := 2 21 | 22 | def sched_default (sc : score) : score := { sc with priority := prio.default } 23 | def sched_now (sc : score) : score := { sc with priority := prio.immediate } 24 | 25 | def inc_cost (sc : score) (n : ℕ) : score := { sc with cost := sc.cost + n } 26 | 27 | def min (a b : score) : score := 28 | { priority := min a.priority b.priority, 29 | in_sos := a.in_sos && b.in_sos, 30 | cost := min a.cost b.cost, 31 | age := min a.age b.age } 32 | 33 | def combine (a b : score) : score := 34 | { priority := max a.priority b.priority, 35 | in_sos := a.in_sos && b.in_sos, 36 | cost := a.cost + b.cost, 37 | age := max a.age b.age } 38 | end score 39 | 40 | namespace score 41 | meta instance : has_to_string score := 42 | ⟨λe, "[" ++ to_string e.priority ++ 43 | "," ++ to_string e.cost ++ 44 | "," ++ to_string e.age ++ 45 | ",sos=" ++ to_string e.in_sos ++ "]"⟩ 46 | end score 47 | 48 | def clause_id := ℕ 49 | namespace clause_id 50 | def to_nat (id : clause_id) : ℕ := id 51 | instance : decidable_eq clause_id := nat.decidable_eq 52 | instance : has_lt clause_id := nat.has_lt 53 | instance : decidable_rel ((<) : clause_id → clause_id → Prop) := nat.decidable_lt 54 | 55 | end clause_id 56 | 57 | meta structure derived_clause := 58 | (id : clause_id) 59 | (c : clause) 60 | (selected : list ℕ) 61 | (assertions : list expr) 62 | (sc : score) 63 | 64 | namespace derived_clause 65 | 66 | meta instance : has_to_tactic_format derived_clause := 67 | ⟨λc, do 68 | prf_fmt ← pp c.c.proof, 69 | c_fmt ← pp c.c, 70 | ass_fmt ← pp (c.assertions.map (λa, a.local_type)), 71 | return $ 72 | to_string c.sc ++ " " ++ 73 | prf_fmt ++ " " ++ 74 | c_fmt ++ " <- " ++ ass_fmt ++ 75 | " (selected: " ++ to_fmt c.selected ++ 76 | ")" 77 | ⟩ 78 | 79 | meta def clause_with_assertions (ac : derived_clause) : clause := 80 | ac.c.close_constn ac.assertions 81 | 82 | meta def update_proof (dc : derived_clause) (p : expr) : derived_clause := 83 | { dc with c := { dc.c with proof := p } } 84 | 85 | end derived_clause 86 | 87 | meta structure locked_clause := 88 | (dc : derived_clause) 89 | (reasons : list (list expr)) 90 | 91 | namespace locked_clause 92 | 93 | meta instance : has_to_tactic_format locked_clause := 94 | ⟨λc, do 95 | c_fmt ← pp c.dc, 96 | reasons_fmt ← pp (c.reasons.map (λr, r.map (λa, a.local_type))), 97 | return $ c_fmt ++ " (locked in case of: " ++ reasons_fmt ++ ")" 98 | ⟩ 99 | 100 | end locked_clause 101 | 102 | meta structure prover_state := 103 | (active : rb_map clause_id derived_clause) 104 | (passive : rb_map clause_id derived_clause) 105 | (newly_derived : list derived_clause) 106 | (prec : list expr) 107 | (locked : list locked_clause) 108 | (local_false : expr) 109 | (sat_solver : cdcl.state) 110 | (current_model : rb_map expr bool) 111 | (sat_hyps : rb_map expr (expr × expr)) 112 | (needs_sat_run : bool) 113 | (clause_counter : nat) 114 | 115 | open prover_state 116 | 117 | private meta def join_with_nl : list format → format := 118 | list.foldl (λx y, x ++ format.line ++ y) format.nil 119 | 120 | private meta def prover_state_tactic_fmt (s : prover_state) : tactic format := do 121 | active_fmts ← mapm pp $ rb_map.values s.active, 122 | passive_fmts ← mapm pp $ rb_map.values s.passive, 123 | new_fmts ← mapm pp s.newly_derived, 124 | locked_fmts ← mapm pp s.locked, 125 | sat_fmts ← mapm pp s.sat_solver.clauses, 126 | sat_model_fmts ← s.current_model.to_list.mmap (λx, if x.2 = tt then pp x.1 else pp `(not %%x.1)), 127 | prec_fmts ← mapm pp s.prec, 128 | return (join_with_nl 129 | ([to_fmt "active:"] ++ ((append (to_fmt " ")) <$> active_fmts) ++ 130 | [to_fmt "passive:"] ++ ((append (to_fmt " ")) <$> passive_fmts) ++ 131 | [to_fmt "new:"] ++ ((append (to_fmt " ")) <$> new_fmts) ++ 132 | [to_fmt "locked:"] ++ ((append (to_fmt " ")) <$> locked_fmts) ++ 133 | [to_fmt "sat formulas:"] ++ ((append (to_fmt " ")) <$> sat_fmts) ++ 134 | [to_fmt "sat model:"] ++ ((append (to_fmt " ")) <$> sat_model_fmts) ++ 135 | [to_fmt "precedence order: " ++ to_fmt prec_fmts])) 136 | 137 | meta instance : has_to_tactic_format prover_state := 138 | ⟨prover_state_tactic_fmt⟩ 139 | 140 | meta def prover := state_t prover_state tactic 141 | 142 | namespace prover 143 | local attribute [reducible] prover cdcl.solver 144 | 145 | meta instance : monad prover := infer_instance 146 | meta instance : alternative prover := infer_instance 147 | meta instance : monad_state _ prover := infer_instance 148 | meta instance : monad_state_adapter _ _ cdcl.solver prover := infer_instance 149 | meta instance : has_monad_lift tactic prover := infer_instance 150 | meta instance (α : Type) : has_coe (tactic α) (prover α) := 151 | ⟨monad_lift⟩ 152 | 153 | end prover 154 | 155 | meta def selection_strategy := derived_clause → prover derived_clause 156 | 157 | meta def get_active : prover (rb_map clause_id derived_clause) := 158 | do state ← get, return state.active 159 | 160 | meta def add_active (a : derived_clause) : prover punit := 161 | do state ← get, 162 | put { state with active := state.active.insert a.id a } 163 | 164 | meta def get_passive : prover (rb_map clause_id derived_clause) := 165 | passive <$> get 166 | 167 | meta def get_precedence : prover (list expr) := 168 | do state ← get, return state.prec 169 | 170 | meta def get_term_order : prover (expr → expr → bool) := do 171 | state ← get, 172 | return $ mk_lpo (name_of_funsym <$> state.prec) 173 | 174 | private meta def set_precedence (new_prec : list expr) : prover punit := 175 | do state ← get, put { state with prec := new_prec } 176 | 177 | meta def register_consts_in_precedence (consts : list expr) := do 178 | p ← get_precedence, 179 | p_set ← return (rb_map.set_of_list (name_of_funsym <$> p)), 180 | new_syms ← return $ list.filter (λc, ¬p_set.contains (name_of_funsym c)) consts, 181 | set_precedence (new_syms ++ p) 182 | 183 | meta def in_sat_solver {A} : cdcl.solver A → prover A := 184 | adapt_state (λ st, (st.sat_solver, st)) (λ sat st, { sat_solver := sat, ..st }) 185 | 186 | meta def collect_ass_hyps (c : clause) : prover (list expr) := 187 | let lcs := contained_lconsts c.proof in 188 | do st ← get, 189 | return (do 190 | hs ← st.sat_hyps.values, 191 | h ← [hs.1, hs.2], 192 | guard $ lcs.contains h.local_uniq_name, 193 | [h]) 194 | 195 | meta def get_clause_count : prover ℕ := 196 | do s ← get, return s.clause_counter 197 | 198 | meta def get_new_cls_id : prover clause_id := do 199 | state ← get, 200 | put { state with clause_counter := state.clause_counter + 1 }, 201 | return state.clause_counter 202 | 203 | meta def mk_derived (c : clause) (sc : score) : prover derived_clause := do 204 | ass ← collect_ass_hyps c, 205 | id ← get_new_cls_id, 206 | return { id := id, c := c, selected := [], assertions := ass, sc := sc } 207 | 208 | meta def add_inferred (c : derived_clause) : prover unit := do 209 | c' ← c.c.normalize, c' ← return { c with c := c' }, 210 | register_consts_in_precedence (contained_funsyms c'.c.type).values, 211 | state ← get, 212 | put { state with newly_derived := c' :: state.newly_derived }, 213 | skip 214 | 215 | 216 | 217 | -- FIXME: what if we've seen the variable before, but with a weaker score? 218 | meta def mk_sat_var (v : expr) (suggested_ph : bool) (suggested_ev : score) : prover unit := 219 | do st ← get, if st.sat_hyps.contains v then return () else do 220 | hpv ← mk_local_def `h v, 221 | hnv ← mk_local_def `hn $ imp v st.local_false, 222 | modify $ λst, { st with sat_hyps := st.sat_hyps.insert v (hpv, hnv) }, 223 | in_sat_solver $ cdcl.mk_var_core v suggested_ph, 224 | match v with 225 | | (pi _ _ _ _) := do 226 | c ← clause.of_proof st.local_false hpv, 227 | mk_derived c suggested_ev >>= add_inferred 228 | | _ := do cp ← clause.of_proof st.local_false hpv, mk_derived cp suggested_ev >>= add_inferred, 229 | cn ← clause.of_proof st.local_false hnv, mk_derived cn suggested_ev >>= add_inferred 230 | end 231 | 232 | meta def get_sat_hyp_core (v : expr) (ph : bool) : prover (option expr) := 233 | flip (<$>) get $ λst, 234 | match st.sat_hyps.find v with 235 | | some (hp, hn) := some $ if ph then hp else hn 236 | | none := none 237 | end 238 | 239 | meta def get_sat_hyp (v : expr) (ph : bool) : prover expr := 240 | do hyp_opt ← get_sat_hyp_core v ph, 241 | match hyp_opt with 242 | | some hyp := return hyp 243 | | none := fail $ "unknown sat variable: " ++ v.to_string 244 | end 245 | 246 | meta def add_sat_clause (c : clause) (suggested_ev : score) : prover unit := do 247 | c ← c.distinct, 248 | already_added ← flip (<$>) get $ λst, decidable.to_bool $ 249 | c.type ∈ st.sat_solver.clauses.map (λd, d.type), 250 | if already_added then return () else do 251 | c.get_lits.mmap' $ λl, mk_sat_var l.formula l.is_neg suggested_ev, 252 | in_sat_solver $ cdcl.mk_clause c, 253 | modify $ λst, { st with needs_sat_run := tt }, 254 | skip 255 | 256 | meta def sat_eval_lit (v : expr) (pol : bool) : prover bool := 257 | do v_st ← flip (<$>) get $ λst, st.current_model.find v, 258 | match v_st with 259 | | some ph := return $ if pol then ph else bnot ph 260 | | none := return tt 261 | end 262 | 263 | meta def sat_eval_assertion (assertion : expr) : prover bool := 264 | do lf ← flip (<$>) get $ λst, st.local_false, 265 | match is_local_not lf assertion.local_type with 266 | | some v := sat_eval_lit v ff 267 | | none := sat_eval_lit assertion.local_type tt 268 | end 269 | 270 | meta def sat_eval_assertions : list expr → prover bool 271 | | (a::ass) := do v_a ← sat_eval_assertion a, 272 | if v_a then 273 | sat_eval_assertions ass 274 | else 275 | return ff 276 | | [] := return tt 277 | 278 | private meta def intern_clause (c : derived_clause) : prover derived_clause := do 279 | hyp_name ← get_unused_name (mk_simple_name $ "clause_" ++ to_string c.id.to_nat) none, 280 | c' ← return $ c.c.close_constn c.assertions, 281 | assertv hyp_name c'.type c'.proof, 282 | proof' ← get_local hyp_name, 283 | type ← infer_type proof', -- FIXME: otherwise "" 284 | return $ c.update_proof $ app_of_list proof' c.assertions 285 | 286 | meta def register_as_passive (c : derived_clause) : prover unit := do 287 | c ← intern_clause c, 288 | ass_v ← sat_eval_assertions c.assertions, 289 | if c.c.num_quants = 0 ∧ c.c.num_lits = 0 then 290 | add_sat_clause c.clause_with_assertions c.sc 291 | else if ¬ass_v then do 292 | modify $ λst, { st with locked := ⟨c, []⟩ :: st.locked }, skip 293 | else do 294 | modify $ λst, { st with passive := st.passive.insert c.id c }, skip 295 | 296 | meta def remove_passive (id : clause_id) : prover unit := 297 | do state ← get, put { state with passive := state.passive.erase id }, skip 298 | 299 | meta def move_locked_to_passive : prover unit := do 300 | locked ← flip (<$>) get (λst, st.locked), 301 | new_locked ← flip filter locked (λlc, do 302 | reason_vals ← mapm sat_eval_assertions lc.reasons, 303 | c_val ← sat_eval_assertions lc.dc.assertions, 304 | if reason_vals.for_all (λr, r = ff) ∧ c_val then do 305 | modify $ λst, { st with passive := st.passive.insert lc.dc.id lc.dc }, 306 | return ff 307 | else 308 | return tt 309 | ), 310 | modify $ λst, { st with locked := new_locked }, skip 311 | 312 | meta def move_active_to_locked : prover unit := 313 | do active ← get_active, active.values.mmap' $ λac, do 314 | c_val ← sat_eval_assertions ac.assertions, 315 | if ¬c_val then do 316 | modify $ λst, { st with 317 | active := st.active.erase ac.id, 318 | locked := ⟨ac, []⟩ :: st.locked 319 | }, skip 320 | else 321 | return () 322 | 323 | meta def move_passive_to_locked : prover unit := 324 | do passive ← flip (<$>) get $ λst, st.passive, passive.to_list.mmap' $ λpc, do 325 | c_val ← sat_eval_assertions pc.2.assertions, 326 | if ¬c_val then do 327 | modify $ λst, { st with 328 | passive := st.passive.erase pc.1, 329 | locked := ⟨pc.2, []⟩ :: st.locked 330 | }, skip 331 | else 332 | return () 333 | 334 | def super_cc_config : cc_config := 335 | { em := ff } 336 | 337 | meta def do_sat_run : prover (option expr) := 338 | do sat_result ← in_sat_solver $ cdcl.run (cdcl.theory_solver_of_tactic failure), 339 | modify $ λst, { st with needs_sat_run := ff }, 340 | old_model ← prover_state.current_model <$> get, 341 | match sat_result with 342 | | (cdcl.result.unsat proof) := return (some proof) 343 | | (cdcl.result.sat new_model) := do 344 | modify $ λst, { st with current_model := new_model }, 345 | move_locked_to_passive, 346 | move_active_to_locked, 347 | move_passive_to_locked, 348 | return none 349 | end 350 | 351 | meta def take_newly_derived : prover (list derived_clause) := do 352 | state ← get, 353 | put { state with newly_derived := [] }, 354 | return state.newly_derived 355 | 356 | meta def remove_redundant (id : clause_id) (parents : list derived_clause) : prover unit := do 357 | when (not $ parents.for_all $ λp, p.id ≠ id) (fail "clause is redundant because of itself"), 358 | red ← flip (<$>) get (λst, st.active.find id), 359 | match red with 360 | | none := return () 361 | | some red := do 362 | let reasons := parents.map (λp, p.assertions), 363 | let assertion := red.assertions, 364 | if reasons.for_all $ λr, r.subset_of assertion then do 365 | modify $ λst, { st with active := st.active.erase id }, skip 366 | else do 367 | modify $ λst, { st with active := st.active.erase id, 368 | locked := ⟨red, reasons⟩ :: st.locked }, skip 369 | end 370 | 371 | meta def inference := derived_clause → prover unit 372 | meta structure inf_decl := (prio : ℕ) (inf : inference) 373 | @[user_attribute] 374 | meta def inf_attr : user_attribute := 375 | {name := `super.inf, descr := "inference for the super prover"} 376 | 377 | meta def seq_inferences : list inference → inference 378 | | [] := λgiven, return () 379 | | (inf::infs) := λgiven, do 380 | inf given, 381 | now_active ← get_active, 382 | if rb_map.contains now_active given.id then 383 | seq_inferences infs given 384 | else 385 | return () 386 | 387 | meta def simp_inference (simpl : derived_clause → prover (option clause)) : inference := 388 | λgiven, do maybe_simpld ← simpl given, 389 | match maybe_simpld with 390 | | some simpld := do 391 | derived_simpld ← mk_derived simpld given.sc.sched_now, 392 | add_inferred derived_simpld, 393 | remove_redundant given.id [] 394 | | none := return () 395 | end 396 | 397 | meta def preprocessing_rule (f : list derived_clause → prover (list derived_clause)) : prover unit := do 398 | state ← get, 399 | newly_derived' ← f state.newly_derived, 400 | state' ← get, 401 | put { state' with newly_derived := newly_derived' }, skip 402 | 403 | meta def clause_selection_strategy := ℕ → prover clause_id 404 | 405 | namespace prover_state 406 | 407 | meta def empty (local_false : expr) : prover_state := 408 | { active := rb_map.mk _ _, passive := rb_map.mk _ _, 409 | newly_derived := [], prec := [], clause_counter := 0, 410 | local_false := local_false, 411 | locked := [], sat_solver := cdcl.state.initial local_false, 412 | current_model := rb_map.mk _ _, sat_hyps := rb_map.mk _ _, needs_sat_run := ff } 413 | 414 | meta def initial (local_false : expr) (clauses : list clause) : tactic prover_state := do 415 | after_setup ← (clauses.mmap' (λc : clause, 416 | let in_sos := ((contained_lconsts c.proof).erase local_false.local_uniq_name).size = 0 in 417 | do mk_derived c { priority := score.prio.immediate, in_sos := in_sos, 418 | age := 0, cost := 0 } >>= add_inferred 419 | )).run (empty local_false), 420 | return after_setup.2 421 | 422 | end prover_state 423 | 424 | meta def inf_score (add_cost : ℕ) (scores : list score) : prover score := do 425 | age ← get_clause_count, 426 | return $ list.foldl score.combine { priority := score.prio.default, 427 | in_sos := tt, 428 | age := age, 429 | cost := add_cost 430 | } scores 431 | 432 | meta def inf_if_successful (add_cost : ℕ) (parent : derived_clause) (tac : tactic (list clause)) : prover unit := 433 | (do inferred ← tac, 434 | inferred.mmap' $ λc, 435 | inf_score add_cost [parent.sc] >>= mk_derived c >>= add_inferred) 436 | <|> return () 437 | 438 | meta def simp_if_successful (parent : derived_clause) (tac : tactic (list clause)) : prover unit := 439 | (do inferred ← tac, 440 | inferred.mmap' $ λc, 441 | mk_derived c parent.sc.sched_now >>= add_inferred, 442 | remove_redundant parent.id []) 443 | <|> return () 444 | 445 | end super 446 | -------------------------------------------------------------------------------- /src/super/resolution.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause .prover_state .utils 7 | open tactic monad native 8 | 9 | namespace super 10 | 11 | variable gt : expr → expr → bool 12 | variables (ac1 ac2 : derived_clause) 13 | variables (c1 c2 : clause) 14 | variables (i1 i2 : nat) 15 | 16 | -- c1 : ... → ¬a → ... 17 | -- c2 : ... → a → ... 18 | meta def try_resolve : tactic clause := do 19 | qf1 ← c1.open_metan c1.num_quants, 20 | qf2 ← c2.open_metan c2.num_quants, 21 | -- FIXME: using default transparency unifies m*n with (x*y*z)*w here ??? 22 | unify (qf1.1.get_lit i1).formula (qf2.1.get_lit i2).formula transparency.reducible, 23 | qf1i ← qf1.1.inst_mvars, 24 | guard $ clause.is_maximal gt qf1i i1, 25 | op1 ← qf1.1.open_constn i1, 26 | op2 ← qf2.1.open_constn c2.num_lits, 27 | a_in_op2 ← (op2.2.nth i2).to_monad, 28 | clause.meta_closure (qf1.2 ++ qf2.2) $ 29 | (op1.1.inst (op2.1.close_const a_in_op2).proof).close_constn (op1.2 ++ op2.2.remove_nth i2) 30 | 31 | meta def try_add_resolvent : prover unit := do 32 | c' ← try_resolve gt ac1.c ac2.c i1 i2, 33 | inf_score 1 [ac1.sc, ac2.sc] >>= mk_derived c' >>= add_inferred 34 | 35 | meta def maybe_add_resolvent : prover unit := 36 | try_add_resolvent gt ac1 ac2 i1 i2 <|> return () 37 | 38 | meta def resolution_left_inf : inference := 39 | assume given, do active ← get_active, sequence' $ do 40 | given_i ← given.selected, 41 | guard $ clause.literal.is_neg (given.c.get_lit given_i), 42 | other ← rb_map.values active, 43 | guard $ ¬given.sc.in_sos ∨ ¬other.sc.in_sos, 44 | other_i ← other.selected, 45 | guard $ clause.literal.is_pos (other.c.get_lit other_i), 46 | [maybe_add_resolvent gt other given other_i given_i] 47 | 48 | meta def resolution_right_inf : inference := 49 | assume given, do active ← get_active, sequence' $ do 50 | given_i ← given.selected, 51 | guard $ clause.literal.is_pos (given.c.get_lit given_i), 52 | other ← rb_map.values active, 53 | guard $ ¬given.sc.in_sos ∨ ¬other.sc.in_sos, 54 | other_i ← other.selected, 55 | guard $ clause.literal.is_neg (other.c.get_lit other_i), 56 | [maybe_add_resolvent gt given other given_i other_i] 57 | 58 | @[super.inf] 59 | meta def resolution_inf : inf_decl := inf_decl.mk 40 $ 60 | assume given, do gt ← get_term_order, resolution_left_inf gt given >> resolution_right_inf gt given 61 | 62 | end super 63 | -------------------------------------------------------------------------------- /src/super/selection.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .prover_state 7 | 8 | open native 9 | 10 | namespace super 11 | 12 | meta def simple_selection_strategy (f : (expr → expr → bool) → clause → list ℕ) : selection_strategy := 13 | assume dc, do gt ← get_term_order, return $ 14 | if dc.selected.empty ∧ dc.c.num_lits > 0 15 | then { dc with selected := f gt dc.c } 16 | else dc 17 | 18 | meta def dumb_selection : selection_strategy := 19 | simple_selection_strategy $ λgt c, 20 | match c.lits_where clause.literal.is_neg with 21 | | [] := list.range c.num_lits 22 | | neg_lit::_ := [neg_lit] 23 | end 24 | 25 | meta def selection21 : selection_strategy := 26 | simple_selection_strategy $ λgt c, 27 | let maximal_lits := list.filter_maximal (λi j, 28 | gt (c.get_lit i).formula (c.get_lit j).formula) (list.range c.num_lits) in 29 | if list.length maximal_lits = 1 then maximal_lits else 30 | let neg_lits := list.filter (λi, (c.get_lit i).is_neg) (list.range c.num_lits), 31 | maximal_neg_lits := list.filter_maximal (λi j, 32 | gt (c.get_lit i).formula (c.get_lit j).formula) neg_lits in 33 | if ¬maximal_neg_lits.empty then 34 | list.take 1 maximal_neg_lits 35 | else 36 | maximal_lits 37 | 38 | meta def selection22 : selection_strategy := 39 | simple_selection_strategy $ λgt c, 40 | let maximal_lits := list.filter_maximal (λi j, 41 | gt (c.get_lit i).formula (c.get_lit j).formula) (list.range c.num_lits), 42 | maximal_lits_neg := list.filter (λi, (c.get_lit i).is_neg) maximal_lits in 43 | if ¬maximal_lits_neg.empty then 44 | list.take 1 maximal_lits_neg 45 | else 46 | maximal_lits 47 | 48 | def sum {α} [has_zero α] [has_add α] : list α → α 49 | | [] := 0 50 | | (x::xs) := x + sum xs 51 | 52 | meta def clause_weight (c : derived_clause) : nat := 53 | sum (c.c.get_lits.map (λl : clause.literal, expr_size l.formula + if l.is_pos then 10 else 1)) 54 | 55 | meta def find_minimal_by (passive : rb_map clause_id derived_clause) 56 | {A} [has_lt A] [decidable_rel ((<) : A → A → Prop)] 57 | (f : derived_clause → A) : clause_id := 58 | match rb_map.min $ rb_map.of_list $ passive.values.map $ λc, (f c, c.id) with 59 | | some id := id 60 | | none := nat.zero 61 | end 62 | 63 | meta def age_of_clause_id : name → ℕ 64 | | (name.mk_numeral i _) := unsigned.to_nat i 65 | | _ := 0 66 | 67 | local attribute [instance] 68 | def prod_has_lt {α β} [has_lt α] [has_lt β] : has_lt (α × β) := 69 | ⟨prod.lex (<) (<)⟩ 70 | 71 | local attribute [instance] 72 | def prod_has_lt_decidable {α β} [has_lt α] [has_lt β] [decidable_rel (@has_lt.lt α _)] [decidable_rel (@has_lt.lt β _)] [decidable_eq α] : 73 | decidable_rel (@has_lt.lt (α × β) _) 74 | | (a,b) (c,d) := if h : a < c then is_true (prod.lex.left _ _ h) else 75 | if h : a = c then 76 | if h2 : b < d then is_true (h ▸ prod.lex.right _ h2) else 77 | is_false (by intro h3; subst h; cases h3; contradiction) 78 | else 79 | is_false (by intro h2; cases h2; contradiction) 80 | 81 | meta def find_minimal_weight (passive : rb_map clause_id derived_clause) : clause_id := 82 | find_minimal_by passive $ λc, (c.sc.priority, clause_weight c + c.sc.cost, c.sc.age, c.id) 83 | 84 | meta def find_minimal_age (passive : rb_map clause_id derived_clause) : clause_id := 85 | find_minimal_by passive $ λc, (c.sc.priority, c.sc.age, c.id) 86 | 87 | meta def weight_clause_selection : clause_selection_strategy := 88 | assume iter, do state ← get, return $ find_minimal_weight state.passive 89 | 90 | meta def oldest_clause_selection : clause_selection_strategy := 91 | assume iter, do state ← get, return $ find_minimal_age state.passive 92 | 93 | meta def age_weight_clause_selection (thr mod : ℕ) : clause_selection_strategy := 94 | assume iter, if iter % mod < thr then 95 | weight_clause_selection iter 96 | else 97 | oldest_clause_selection iter 98 | 99 | end super 100 | -------------------------------------------------------------------------------- /src/super/simp.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause_ops .prover_state 7 | open tactic monad 8 | 9 | namespace super 10 | 11 | meta def prove_using_assumption : tactic unit := do 12 | tgt ← target, 13 | ass ← mk_local_def `h tgt, 14 | exact ass 15 | 16 | meta def simplify_capturing_assumptions (type : expr) : tactic (expr × expr × list expr) := do 17 | S ← simp_lemmas.mk_default, 18 | (type', heq, _) ← simplify S [] type, 19 | hyps ← return $ contained_lconsts type, 20 | hyps' ← return $ contained_lconsts_list [type', heq], 21 | add_hyps ← return $ list.filter (λn : expr, ¬hyps.contains n.local_uniq_name) hyps'.values, 22 | return (type', heq, add_hyps) 23 | 24 | meta def try_simplify_left (c : clause) (i : ℕ) : tactic (list clause) := 25 | on_left_at c i $ λtype, do 26 | (type', heq, add_hyps) ← simplify_capturing_assumptions type, 27 | hyp ← mk_local_def `h type', 28 | prf ← mk_eq_mpr heq hyp, 29 | return [(hyp::add_hyps, prf)] 30 | 31 | meta def try_simplify_right (c : clause) (i : ℕ) : tactic (list clause) := 32 | on_right_at' c i $ λhyp, do 33 | (type', heq, add_hyps) ← simplify_capturing_assumptions hyp.local_type, 34 | heqtype ← infer_type heq, 35 | heqsymm ← mk_eq_symm heq, 36 | prf ← mk_eq_mpr heqsymm hyp, 37 | return [(add_hyps, prf)] 38 | 39 | meta def simp_inf : inf_decl := inf_decl.mk 40 $ assume given, sequence' $ do 40 | r ← [try_simplify_right, try_simplify_left], 41 | i ← list.range given.c.num_lits, 42 | [inf_if_successful 2 given (r given.c i)] 43 | 44 | end super 45 | -------------------------------------------------------------------------------- /src/super/splitting.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .prover_state 7 | open monad expr list tactic 8 | 9 | namespace super 10 | 11 | private meta def find_components : list expr → list (list (expr × ℕ)) → list (list (expr × ℕ)) 12 | | (e::es) comps := 13 | let (contain_e, do_not_contain_e) := 14 | partition (λc : list (expr × ℕ), c.exists_ $ λf, 15 | (abstract_local f.1 e.local_uniq_name).has_var) comps in 16 | find_components es $ list.join contain_e :: do_not_contain_e 17 | | _ comps := comps 18 | 19 | meta def get_components (hs : list expr) : list (list expr) := 20 | (find_components hs (hs.zip_with_index.map $ λh, [h])).map $ λc, 21 | (sort_on (λh : expr × ℕ, h.2) c).map $ λh, h.1 22 | 23 | meta def extract_assertions : clause → prover (clause × list expr) | c := 24 | if c.num_lits = 0 then return (c, []) 25 | else if c.num_quants ≠ 0 then do 26 | qf ← c.open_constn c.num_quants, 27 | qf_wo_as ← extract_assertions qf.1, 28 | return (qf_wo_as.1.close_constn qf.2, qf_wo_as.2) 29 | else do 30 | hd ← return $ c.get_lit 0, 31 | hyp_opt ← get_sat_hyp_core hd.formula hd.is_neg, 32 | match hyp_opt with 33 | | some h := do 34 | wo_as ← extract_assertions (c.inst h), 35 | return (wo_as.1, h :: wo_as.2) 36 | | _ := do 37 | op ← c.open_const, 38 | op_wo_as ← extract_assertions op.1, 39 | return (op_wo_as.1.close_const op.2, op_wo_as.2) 40 | end 41 | 42 | meta def mk_splitting_clause' (empty_clause : clause) : list (list expr) → tactic (list expr × expr) 43 | | [] := return ([], empty_clause.proof) 44 | | ([p] :: comps) := do p' ← mk_splitting_clause' comps, return (p::p'.1, p'.2) 45 | | (comp :: comps) := do 46 | (hs, p') ← mk_splitting_clause' comps, 47 | hnc ← mk_local_def `hnc (imp (pis comp empty_clause.local_false) empty_clause.local_false), 48 | p'' ← return $ app hnc (lambdas comp p'), 49 | return (hnc::hs, p'') 50 | 51 | meta def mk_splitting_clause (empty_clause : clause) (comps : list (list expr)) : tactic clause := do 52 | (hs, p) ← mk_splitting_clause' empty_clause comps, 53 | return $ { empty_clause with proof := p }.close_constn hs 54 | 55 | @[super.inf] 56 | meta def splitting_inf : inf_decl := inf_decl.mk 30 $ assume given, do 57 | lf ← flip (<$>) get $ λst, st.local_false, 58 | op ← given.c.open_constn given.c.num_binders, 59 | if list.bor (given.c.get_lits.map $ λl, (is_local_not lf l.formula).is_some) then return () else 60 | let comps := get_components op.2 in 61 | if comps.length < 2 then return () else do 62 | splitting_clause ← mk_splitting_clause op.1 comps, 63 | ass ← collect_ass_hyps splitting_clause, 64 | add_sat_clause (splitting_clause.close_constn ass) given.sc.sched_default, 65 | remove_redundant given.id [] 66 | 67 | end super 68 | -------------------------------------------------------------------------------- /src/super/subsumption.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause .prover_state 7 | open tactic monad native 8 | 9 | namespace super 10 | 11 | private meta def try_subsume_core : list clause.literal → list clause.literal → tactic unit 12 | | [] _ := skip 13 | | small large := first $ do 14 | i ← small.zip_with_index, j ← large.zip_with_index, 15 | return $ do 16 | unify_lit i.1 j.1, 17 | try_subsume_core (small.remove_nth i.2) (large.remove_nth j.2) 18 | 19 | -- FIXME: this is incorrect if a quantifier is unused 20 | meta def try_subsume (small large : clause) : tactic unit := do 21 | small_open ← clause.open_metan small (clause.num_quants small), 22 | large_open ← clause.open_constn large (clause.num_quants large), 23 | guard $ small.num_lits ≤ large.num_lits, 24 | try_subsume_core small_open.1.get_lits large_open.1.get_lits 25 | 26 | meta def does_subsume (small large : clause) : tactic bool := 27 | (try_subsume small large >> return tt) <|> return ff 28 | 29 | meta def does_subsume_with_assertions (small large : derived_clause) : prover bool := do 30 | if small.assertions.subset_of large.assertions then do 31 | does_subsume small.c large.c 32 | else 33 | return ff 34 | 35 | meta def any_tt {m : Type → Type} [monad m] (active : rb_map clause_id derived_clause) (pred : derived_clause → m bool) : m bool := 36 | active.fold (return ff) $ λk a cont, do 37 | v ← pred a, if v then return tt else cont 38 | 39 | meta def any_tt_list {m : Type → Type} [monad m] {A} (pred : A → m bool) : list A → m bool 40 | | [] := return ff 41 | | (x::xs) := do v ← pred x, if v then return tt else any_tt_list xs 42 | 43 | @[super.inf] 44 | meta def forward_subsumption : inf_decl := inf_decl.mk 20 $ 45 | assume given, do active ← get_active, 46 | sequence' $ do a ← active.values, 47 | guard $ a.id ≠ given.id, 48 | return $ do 49 | ss ← does_subsume a.c given.c, 50 | if ss 51 | then remove_redundant given.id [a] 52 | else return () 53 | 54 | meta def forward_subsumption_pre : prover unit := preprocessing_rule $ λnew, do 55 | active ← get_active, filter (λn, do 56 | do ss ← any_tt active (λa, 57 | if a.assertions.subset_of n.assertions then do 58 | does_subsume a.c n.c 59 | else 60 | -- TODO: move to locked 61 | return ff), 62 | return (bnot ss)) new 63 | 64 | meta def subsumption_interreduction : list derived_clause → prover (list derived_clause) 65 | | (c::cs) := do 66 | -- TODO: move to locked 67 | cs_that_subsume_c ← filter (λd, does_subsume_with_assertions d c) cs, 68 | if ¬cs_that_subsume_c.empty then 69 | -- TODO: update score 70 | subsumption_interreduction cs 71 | else do 72 | cs_not_subsumed_by_c ← filter (λd, bnot <$> does_subsume_with_assertions c d) cs, 73 | cs' ← subsumption_interreduction cs_not_subsumed_by_c, 74 | return (c::cs') 75 | | [] := return [] 76 | 77 | meta def subsumption_interreduction_pre : prover unit := 78 | preprocessing_rule $ λnew, 79 | let new' := list.sort_on (λc : derived_clause, c.c.num_lits) new in 80 | subsumption_interreduction new' 81 | 82 | meta def keys_where_tt {m} {K V : Type} [monad m] (active : rb_map K V) (pred : V → m bool) : m (list K) := 83 | @rb_map.fold _ _ (m (list K)) active (return []) $ λk a cont, do 84 | v ← pred a, rest ← cont, return $ if v then k::rest else rest 85 | 86 | @[super.inf] 87 | meta def backward_subsumption : inf_decl := inf_decl.mk 20 $ λgiven, do 88 | active ← get_active, 89 | ss ← keys_where_tt active (λa, does_subsume given.c a.c), 90 | sequence' $ do id ← ss, guard (id ≠ given.id), [remove_redundant id [given]] 91 | 92 | end super 93 | -------------------------------------------------------------------------------- /src/super/superposition.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .clause .prover_state .utils 7 | open tactic monad expr native 8 | 9 | namespace super 10 | 11 | def position := list ℕ 12 | 13 | meta def get_rwr_positions : expr → list position 14 | | (app a b) := [[]] ++ 15 | do arg ← list.zip_with_index (get_app_args (app a b)), 16 | pos ← get_rwr_positions arg.1, 17 | [arg.2 :: pos] 18 | | (var _) := [] 19 | | e := [[]] 20 | 21 | meta def get_position : expr → position → expr 22 | | (app a b) (p::ps) := 23 | match list.nth (get_app_args (app a b)) p with 24 | | some arg := get_position arg ps 25 | | none := (app a b) 26 | end 27 | | e _ := e 28 | 29 | meta def replace_position (v : expr) : expr → position → expr 30 | | (app a b) (p::ps) := 31 | let args := get_app_args (app a b) in 32 | match args.nth p with 33 | | some arg := app_of_list a.get_app_fn $ args.update_nth p $ replace_position arg ps 34 | | none := app a b 35 | end 36 | | e [] := v 37 | | e _ := e 38 | 39 | variable gt : expr → expr → bool 40 | variables (c1 c2 : clause) 41 | variables (ac1 ac2 : derived_clause) 42 | variables (i1 i2 : nat) 43 | variable pos : list ℕ 44 | variable ltr : bool 45 | variable lt_in_termorder : bool 46 | variable congr_ax : name 47 | 48 | lemma {u v w} sup_ltr (F : Sort u) (A : Sort v) (a1 a2) (f : A → Sort w) : (f a1 → F) → f a2 → a1 = a2 → F := 49 | assume hnfa1 hfa2 he, hnfa1 (@eq.rec A a2 f hfa2 a1 he.symm) 50 | lemma {u v w} sup_rtl (F : Sort u) (A : Sort v) (a1 a2) (f : A → Sort w) : (f a1 → F) → f a2 → a2 = a1 → F := 51 | assume hnfa1 hfa2 heq, hnfa1 (@eq.rec A a2 f hfa2 a1 heq) 52 | 53 | meta def is_eq_dir (e : expr) (ltr : bool) : option (expr × expr) := 54 | match is_eq e with 55 | | some (lhs, rhs) := if ltr then some (lhs, rhs) else some (rhs, lhs) 56 | | none := none 57 | end 58 | 59 | meta def try_sup : tactic clause := do 60 | guard $ (c1.get_lit i1).is_pos, 61 | qf1 ← c1.open_metan c1.num_quants, 62 | qf2 ← c2.open_metan c2.num_quants, 63 | (rwr_from, rwr_to) ← (is_eq_dir (qf1.1.get_lit i1).formula ltr).to_monad, 64 | atom ← return (qf2.1.get_lit i2).formula, 65 | eq_type ← infer_type rwr_from, 66 | atom_at_pos ← return $ get_position atom pos, 67 | atom_at_pos_type ← infer_type atom_at_pos, 68 | unify eq_type atom_at_pos_type, 69 | unify rwr_from atom_at_pos transparency.none, 70 | rwr_from' ← instantiate_mvars atom_at_pos, 71 | rwr_to' ← instantiate_mvars rwr_to, 72 | if lt_in_termorder 73 | then guard (gt rwr_from' rwr_to') 74 | else guard (¬gt rwr_to' rwr_from'), 75 | rwr_ctx_varn ← mk_fresh_name, 76 | abs_rwr_ctx ← return $ 77 | lam rwr_ctx_varn binder_info.default eq_type 78 | (if (qf2.1.get_lit i2).is_neg 79 | then replace_position (mk_var 0) atom pos 80 | else imp (replace_position (mk_var 0) atom pos) c2.local_false), 81 | lf_univ ← infer_univ c1.local_false, 82 | univ ← infer_univ eq_type, 83 | atom_univ ← infer_univ atom, 84 | op1 ← qf1.1.open_constn i1, 85 | op2 ← qf2.1.open_constn c2.num_lits, 86 | hi2 ← (op2.2.nth i2).to_monad, 87 | new_atom ← whnf_no_delta $ app abs_rwr_ctx rwr_to', 88 | new_hi2 ← return $ local_const hi2.local_uniq_name `H binder_info.default new_atom, 89 | new_fin_prf ← 90 | return $ app_of_list (const congr_ax [lf_univ, univ, atom_univ]) [c1.local_false, eq_type, rwr_from, rwr_to, 91 | abs_rwr_ctx, (op2.1.close_const hi2).proof, new_hi2], 92 | clause.meta_closure (qf1.2 ++ qf2.2) $ (op1.1.inst new_fin_prf).close_constn (op1.2 ++ op2.2.update_nth i2 new_hi2) 93 | 94 | meta def rwr_positions (c : clause) (i : nat) : list (list ℕ) := 95 | get_rwr_positions (c.get_lit i).formula 96 | 97 | meta def try_add_sup : prover unit := 98 | (do c' ← try_sup gt ac1.c ac2.c i1 i2 pos ltr ff congr_ax, 99 | inf_score 2 [ac1.sc, ac2.sc] >>= mk_derived c' >>= add_inferred) 100 | <|> return () 101 | 102 | meta def superposition_back_inf : inference := 103 | assume given, do active ← get_active, sequence' $ do 104 | given_i ← given.selected, 105 | guard (given.c.get_lit given_i).is_pos, 106 | option.to_monad $ is_eq (given.c.get_lit given_i).formula, 107 | other ← rb_map.values active, 108 | guard $ ¬given.sc.in_sos ∨ ¬other.sc.in_sos, 109 | other_i ← other.selected, 110 | pos ← rwr_positions other.c other_i, 111 | -- FIXME(gabriel): ``sup_ltr fails to resolve at runtime 112 | [do try_add_sup gt given other given_i other_i pos tt ``super.sup_ltr, 113 | try_add_sup gt given other given_i other_i pos ff ``super.sup_rtl] 114 | 115 | meta def superposition_fwd_inf : inference := 116 | assume given, do active ← get_active, sequence' $ do 117 | given_i ← given.selected, 118 | other ← rb_map.values active, 119 | guard $ ¬given.sc.in_sos ∨ ¬other.sc.in_sos, 120 | other_i ← other.selected, 121 | guard (other.c.get_lit other_i).is_pos, 122 | option.to_monad $ is_eq (other.c.get_lit other_i).formula, 123 | pos ← rwr_positions given.c given_i, 124 | [do try_add_sup gt other given other_i given_i pos tt ``super.sup_ltr, 125 | try_add_sup gt other given other_i given_i pos ff ``super.sup_rtl] 126 | 127 | @[super.inf] 128 | meta def superposition_inf : inf_decl := inf_decl.mk 40 $ 129 | assume given, do gt ← get_term_order, 130 | superposition_fwd_inf gt given, 131 | superposition_back_inf gt given 132 | 133 | end super 134 | -------------------------------------------------------------------------------- /src/super/trim.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | import .utils 7 | open monad expr tactic 8 | 9 | namespace super 10 | 11 | -- TODO(gabriel): rewrite using conversions 12 | meta def trim : expr → tactic expr 13 | | (app (lam n m d b) arg) := 14 | if ¬b.has_var then 15 | trim b 16 | else 17 | app <$> trim (lam n m d b) <*> trim arg 18 | | (app a b) := app <$> trim a <*> trim b 19 | | (lam n m d b) := do 20 | x ← mk_local' `x m d, 21 | b' ← trim (instantiate_var b x), 22 | return $ lam n m d (abstract_local b' x.local_uniq_name) 23 | | (elet n t v b) := 24 | if has_var b then do 25 | x ← mk_local_def `x t, 26 | b' ← trim (instantiate_var b x), 27 | return $ elet n t v (abstract_local b' x.local_uniq_name) 28 | else 29 | trim b 30 | | e := return e 31 | 32 | -- iterate trim until convergence 33 | meta def trim' : expr → tactic expr 34 | | e := do e' ← trim e, 35 | if e =ₐ e' then 36 | return e 37 | else 38 | trim' e' 39 | 40 | open tactic 41 | 42 | meta def with_trim {α} (tac : tactic α) : tactic α := do 43 | gs ← get_goals, 44 | match gs with 45 | | (g::gs) := do 46 | g' ← infer_type g >>= mk_meta_var, 47 | set_goals [g'], 48 | r ← tac, 49 | done, 50 | set_goals (g::gs), 51 | instantiate_mvars g' >>= trim' >>= exact, 52 | return r 53 | | [] := fail "no goal" 54 | end 55 | 56 | end super 57 | -------------------------------------------------------------------------------- /src/super/utils.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2017 Gabriel Ebner. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Gabriel Ebner 5 | -/ 6 | open tactic expr list native 7 | 8 | meta def get_metas : expr → list expr 9 | | (var _) := [] 10 | | (sort _) := [] 11 | | (const _ _) := [] 12 | | (mvar n pp_n t) := expr.mvar pp_n n t :: get_metas t 13 | | (local_const _ _ _ t) := get_metas t 14 | | (app a b) := get_metas a ++ get_metas b 15 | | (lam _ _ d b) := get_metas d ++ get_metas b 16 | | (pi _ _ d b) := get_metas d ++ get_metas b 17 | | (elet _ t v b) := get_metas t ++ get_metas v ++ get_metas b 18 | | (macro _ _) := [] 19 | 20 | meta def get_meta_type : expr → expr 21 | | (mvar _ _ t) := t 22 | | _ := mk_var 0 23 | 24 | -- TODO(gabriel): think about how to handle the avalanche of implicit arguments 25 | meta def expr_size : expr → nat 26 | | (var _) := 1 27 | | (sort _) := 1 28 | | (const _ _) := 1 29 | | (mvar n pp_n t) := 1 30 | | (local_const _ _ _ _) := 1 31 | | (app a b) := expr_size a + expr_size b 32 | | (lam _ _ d b) := 1 + expr_size b 33 | | (pi _ _ d b) := 1 + expr_size b 34 | | (elet _ t v b) := 1 + expr_size v + expr_size b 35 | | (macro _ _) := 1 36 | 37 | namespace list 38 | 39 | meta def nub {A} [has_lt A] [decidable_rel ((<) : A → A → Prop)] (l : list A) : list A := 40 | rb_map.keys (rb_map.set_of_list l) 41 | 42 | meta def nub_on {A B} [has_lt B] [decidable_rel ((<) : B → B → Prop)] (f : A → B) (l : list A) : list A := 43 | rb_map.values (rb_map.of_list (map (λx, (f x, x)) l)) 44 | 45 | def nub_on' {A B} [decidable_eq B] (f : A → B) : list A → list A 46 | | [] := [] 47 | | (x::xs) := x :: filter (λy, f x ≠ f y) (nub_on' xs) 48 | 49 | def for_all {A} (l : list A) (p : A → Prop) [decidable_pred p] : bool := 50 | list.all l (λx, to_bool (p x)) 51 | 52 | def exists_ {A} (l : list A) (p : A → Prop) [decidable_pred p] : bool := 53 | list.any l (λx, to_bool (p x)) 54 | 55 | def subset_of {A} [decidable_eq A] (xs ys : list A) := 56 | xs.for_all (λx, x ∈ ys) 57 | 58 | def filter_maximal {A} (gt : A → A → bool) (l : list A) : list A := 59 | filter (λx, l.for_all (λy, ¬gt y x)) l 60 | 61 | private def zip_with_index' {A} : ℕ → list A → list (A × ℕ) 62 | | _ nil := nil 63 | | i (x::xs) := (x,i) :: zip_with_index' (i+1) xs 64 | 65 | def zip_with_index {A} : list A → list (A × ℕ) := 66 | zip_with_index' 0 67 | 68 | def sort {A} [has_lt A] [decidable_rel ((<) : A → A → Prop)] : list A → list A := 69 | qsort (λ x y, x < y) 70 | 71 | def sort_on {A B} (f : A → B) [has_lt B] [decidable_rel ((<) : B → B → Prop)] : list A → list A := 72 | qsort (λ x y, f x < f y) 73 | 74 | end list 75 | 76 | meta def name_of_funsym : expr → name 77 | | (local_const uniq _ _ _) := uniq 78 | | (const n _) := n 79 | | _ := name.anonymous 80 | 81 | private meta def contained_funsyms' : expr → rb_map name expr → rb_map name expr 82 | | (var _) m := m 83 | | (sort _) m := m 84 | | (const n ls) m := rb_map.insert m n (const n ls) 85 | | (mvar _ _ t) m := contained_funsyms' t m 86 | | (local_const uniq pp bi t) m := rb_map.insert m uniq (local_const uniq pp bi t) 87 | | (app a b) m := contained_funsyms' a (contained_funsyms' b m) 88 | | (lam _ _ d b) m := contained_funsyms' d (contained_funsyms' b m) 89 | | (pi _ _ d b) m := contained_funsyms' d (contained_funsyms' b m) 90 | | (elet _ t v b) m := contained_funsyms' t (contained_funsyms' v (contained_funsyms' b m)) 91 | | (macro _ _) m := m 92 | 93 | meta def contained_funsyms (e : expr) : rb_map name expr := 94 | contained_funsyms' e (rb_map.mk name expr) 95 | 96 | private meta def contained_lconsts' : expr → rb_map name expr → rb_map name expr 97 | | (var _) m := m 98 | | (sort _) m := m 99 | | (const _ _) m := m 100 | | (mvar _ _ t) m := contained_lconsts' t m 101 | | (local_const uniq pp bi t) m := contained_lconsts' t (rb_map.insert m uniq (local_const uniq pp bi t)) 102 | | (app a b) m := contained_lconsts' a (contained_lconsts' b m) 103 | | (lam _ _ d b) m := contained_lconsts' d (contained_lconsts' b m) 104 | | (pi _ _ d b) m := contained_lconsts' d (contained_lconsts' b m) 105 | | (elet _ t v b) m := contained_lconsts' t (contained_lconsts' v (contained_lconsts' b m)) 106 | | (macro _ _) m := m 107 | 108 | meta def contained_lconsts (e : expr) : rb_map name expr := 109 | contained_lconsts' e (rb_map.mk name expr) 110 | 111 | meta def contained_lconsts_list (es : list expr) : rb_map name expr := 112 | es.foldl (λlcs e, contained_lconsts' e lcs) (rb_map.mk name expr) 113 | 114 | namespace tactic 115 | 116 | meta def infer_univ (type : expr) : tactic level := 117 | do sort_of_type ← infer_type type >>= whnf, 118 | match sort_of_type with 119 | | sort lvl := return lvl 120 | | not_sort := do fmt ← pp not_sort, 121 | fail $ to_fmt "cannot get universe level of sort: " ++ fmt 122 | end 123 | 124 | end tactic 125 | 126 | -------------------------------------------------------------------------------- /test/cdcl_examples.lean: -------------------------------------------------------------------------------- 1 | import super.cdcl 2 | 3 | example {a} : a → ¬a → false := by cdcl 4 | example {a} : a ∨ ¬a := by cdcl 5 | example {a b} : a → (a → b) → b := by cdcl 6 | example {a b c} : (a → b) → (¬a → b) → (b → c) → b ∧ c := by cdcl 7 | 8 | open tactic 9 | 10 | private meta def lit_unification : tactic unit := 11 | do ls ← local_context, first $ do l ← ls, [do apply l, assumption] 12 | 13 | example {p : ℕ → Prop} : p 2 ∨ p 4 → (p (2*2) → p (2+0)) → p (1+1) := 14 | by cdcl_t lit_unification 15 | 16 | example {p : ℕ → Prop} : 17 | list.foldl (λf v, f ∧ (v ∨ ¬v)) true (p <$> list.range 5) := 18 | begin (target >>= whnf >>= change), cdcl end 19 | 20 | example {a b c : Prop} : (a → b) → (b → c) → (a → c) := by cdcl 21 | -------------------------------------------------------------------------------- /test/super_examples.lean: -------------------------------------------------------------------------------- 1 | import super 2 | open tactic 3 | 4 | def prime (n : ℕ) := ∀ d, d ∣ n → d = 1 ∨ d = n 5 | 6 | lemma nat_mul_cancel_one {m n : ℕ} : m ≠ 0 → m * n = m → n = 1 := 7 | by cases m; super nat.zero_lt_succ nat.eq_of_mul_eq_mul_left nat.mul_one 8 | 9 | lemma not_prime_zero : ¬ prime 0 := 10 | by intro h; cases h 2 ⟨0, rfl⟩; cases h_1 11 | 12 | @[simp] lemma nat.dvd_refl (m : ℕ) : m ∣ m := ⟨1, by simp [nat.mul_one]⟩ 13 | @[simp] theorem nat.dvd_mul_left (a b : ℕ) : a ∣ b * a := ⟨b, nat.mul_comm _ _⟩ 14 | 15 | example {m n : ℕ} : prime (m * n) → m = 1 ∨ n = 1 := 16 | by super with prime nat.dvd_refl nat.dvd_mul_right nat.dvd_mul_left 17 | nat_mul_cancel_one not_prime_zero nat.mul_zero nat.zero_mul 18 | 19 | example : nat.zero ≠ nat.succ nat.zero := by super 20 | example (x y : ℕ) : nat.succ x = nat.succ y → x = y := by super 21 | example (i) (a b c : i) : [a,b,c] = [b,c,a] -> a = b ∧ b = c := by super 22 | 23 | definition is_positive (n : ℕ) := n > 0 24 | example (n : ℕ) : n > 0 ↔ is_positive n := by super with is_positive 25 | 26 | example (m n : ℕ) : 0 + m = 0 + n → m = n := 27 | by super with nat.zero_add 28 | 29 | example : ∀x y : ℕ, x + y = y + x := 30 | begin intros, have h : nat.zero = 0 := rfl, induction x, 31 | super with nat.add_zero nat.zero_add, 32 | super with nat.add_succ nat.succ_add end 33 | 34 | example (i) [inhabited i] : nonempty i := by super 35 | example (i) [nonempty i] : ¬(inhabited i → false) := by super 36 | 37 | example : nonempty ℕ := by super 38 | example : ¬(inhabited ℕ → false) := by super 39 | 40 | example {a b} : ¬(b ∨ ¬a) ∨ (a → b) := by super 41 | example {a} : a ∨ ¬a := by super 42 | example {a} : (a ∧ a) ∨ (¬a ∧ ¬a) := by super 43 | example (i) (c : i) (p : i → Prop) (f : i → i) : 44 | p c → (∀x, p x → p (f x)) → p (f (f (f c))) := by super 45 | 46 | example (i) (p : i → Prop) : ∀x, p x → ∃x, p x := by super 47 | 48 | example (i) [nonempty i] (p : i → i → Prop) : (∀x y, p x y) → ∃x, ∀z, p x z := by super 49 | 50 | example (i) [nonempty i] (p : i → Prop) : (∀x, p x) → ¬¬∀x, p x := by super 51 | 52 | -- Requires non-empty domain. 53 | example {i} [nonempty i] (p : i → Prop) : 54 | (∀x y, p x ∨ p y) → ∃x y, p x ∧ p y := by super 55 | 56 | example (i) (a b : i) (p : i → Prop) (H : a = b) : p b → p a := 57 | by super 58 | 59 | example (i) (a b : i) (p : i → Prop) (H : a = b) : p a → p b := 60 | by super 61 | 62 | example (i) (a b : i) (p : i → Prop) (H : a = b) : p b = p a := 63 | by super 64 | 65 | example (i) (c : i) (p : i → Prop) (f g : i → i) : 66 | p c → (∀x, p x → p (f x)) → (∀x, p x → f x = g x) → f (f c) = g (g c) := 67 | by super 68 | 69 | example (i) (p q : i → i → Prop) (a b c d : i) : 70 | (∀x y z, p x y ∧ p y z → p x z) → 71 | (∀x y z, q x y ∧ q y z → q x z) → 72 | (∀x y, q x y → q y x) → 73 | (∀x y, p x y ∨ q x y) → 74 | p a b ∨ q c d := 75 | by super 76 | 77 | -- This example from Davis-Putnam actually requires a non-empty domain 78 | 79 | example (i) [nonempty i] (f g : i → i → Prop) : 80 | ∃x y, ∀z, (f x y → f y z ∧ f z z) ∧ (f x y ∧ g x y → g x z ∧ g z z) := 81 | by super 82 | 83 | example (person) [nonempty person] (drinks : person → Prop) : 84 | ∃canary, drinks canary → ∀other, drinks other := by super 85 | 86 | example {p q : ℕ → Prop} {r} : (∀x y, p x ∧ q y ∧ r) -> ∀x, (p x ∧ r ∧ q x) := by super 87 | -------------------------------------------------------------------------------- /test/super_tests.lean: -------------------------------------------------------------------------------- 1 | import super 2 | 3 | section 4 | open super tactic 5 | example (i : Type) (a b : i) (p : i → Prop) (H : a = b) (Hpa : p a) : true := by do 6 | H ← get_local `H >>= clause.of_classical_proof, 7 | Hpa ← get_local `Hpa >>= clause.of_classical_proof, 8 | a ← get_local `a, 9 | try_sup (λx y, ff) H Hpa 0 0 [0] tt ff ``super.sup_ltr >>= clause.validate, 10 | to_expr ``(trivial) >>= apply 11 | 12 | example (i : Type) (a b : i) (p : i → Prop) (H : a = b) (Hpa : p a → false) (Hpb : p b → false) : true := by do 13 | H ← get_local `H >>= clause.of_classical_proof, 14 | Hpa ← get_local `Hpa >>= clause.of_classical_proof, 15 | Hpb ← get_local `Hpb >>= clause.of_classical_proof, 16 | try_sup (λx y, ff) H Hpa 0 0 [0] tt ff ``super.sup_ltr >>= clause.validate, 17 | try_sup (λx y, ff) H Hpb 0 0 [0] ff ff ``super.sup_rtl >>= clause.validate, 18 | to_expr ``(trivial) >>= apply 19 | 20 | example (i : Type) (p q : i → Prop) (H : ∀x y, p x → q y → false) : true := by do 21 | h ← get_local `H >>= clause.of_classical_proof, 22 | (op, lcs) ← h^.open_constn h^.num_binders, 23 | guard $ (get_components lcs)^.length = 2, 24 | triv 25 | 26 | example (i : Type) (p : i → i → Prop) (H : ∀x y z, p x y → p y z → false) : true := by do 27 | h ← get_local `H >>= clause.of_classical_proof, 28 | (op, lcs) ← h^.open_constn h^.num_binders, 29 | guard $ (get_components lcs)^.length = 1, 30 | triv 31 | 32 | example (i : Type) (p : i → i → Type) (c : i) (h : ∀ (x : i), p x c → p x c) : true := by do 33 | h ← get_local `h, hcls ← clause.of_classical_proof h, 34 | taut ← is_taut hcls, 35 | when (¬taut) failed, 36 | to_expr ``(trivial) >>= apply 37 | 38 | open tactic 39 | example (m n : ℕ) : true := by do 40 | e₁ ← to_expr ```((0 + (m : ℕ)) + 0), 41 | e₂ ← to_expr ```(0 + (0 + (m : ℕ))), 42 | e₃ ← to_expr ```(0 + (m : ℕ)), 43 | prec ← return (contained_funsyms e₁)^.keys, 44 | prec_gt ← return $ prec_gt_of_name_list prec, 45 | guard $ lpo prec_gt e₁ e₃, 46 | guard $ lpo prec_gt e₂ e₃, 47 | to_expr ``(trivial) >>= apply 48 | 49 | /- 50 | open tactic 51 | example (i : Type) (f : i → i) (c d x : i) : true := by do 52 | ef ← get_local `f, ec ← get_local `c, ed ← get_local `d, 53 | syms ← return [ef,ec,ed], 54 | prec_gt ← return $ prec_gt_of_name_list (list.map local_uniq_name [ef, ec, ed]), 55 | sequence' (do s1 ← syms, s2 ← syms, return (do 56 | s1_fmt ← pp s1, s2_fmt ← pp s2, 57 | trace (s1_fmt ++ to_fmt " > " ++ s2_fmt ++ to_fmt ": " ++ to_fmt (prec_gt s1 s2)) 58 | )), 59 | 60 | exprs ← @mapM tactic _ _ _ to_expr [`(f c), `(f (f c)), `(f d), `(f x), `(f (f x))], 61 | sequence' (do e1 ← exprs, e2 ← exprs, return (do 62 | e1_fmt ← pp e1, e2_fmt ← pp e2, 63 | trace (e1_fmt ++ to_fmt" > " ++ e2_fmt ++ to_fmt": " ++ to_fmt (lpo prec_gt e1 e2)) 64 | )), 65 | 66 | mk_const ``true.intro >>= apply 67 | -/ 68 | open monad 69 | example (x y : ℕ) (h : nat.zero = nat.succ nat.zero) (h2 : nat.succ x = nat.succ y) : true := by do 70 | h ← get_local `h >>= clause.of_classical_proof, 71 | h2 ← get_local `h2 >>= clause.of_classical_proof, 72 | cs ← try_no_confusion_eq_r h 0, 73 | cs.mmap' clause.validate, 74 | cs ← try_no_confusion_eq_r h2 0, 75 | cs.mmap' clause.validate, 76 | to_expr ``(trivial) >>= exact 77 | end 78 | --------------------------------------------------------------------------------