Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 06f4778

Browse files
cipher1024digama0
authored andcommitted
feat(tactic/tauto): handle or in goal
1 parent b4a8548 commit 06f4778

File tree

2 files changed

+49
-2
lines changed

2 files changed

+49
-2
lines changed

tactic/interactive.lean

Lines changed: 45 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -301,6 +301,46 @@ meta structure by_elim_opt :=
301301
meta def solve_by_elim (opt : by_elim_opt := { }) : tactic unit :=
302302
solve_by_elim_aux opt.discharger opt.restr_hyp_set opt.max_rep
303303

304+
theorem or_iff_not_imp_left {a b} [decidable a] : a ∨ b ↔ (¬ a → b) :=
305+
⟨or.resolve_left, λ h, dite _ or.inl (or.inr ∘ h)⟩
306+
307+
theorem or_iff_not_imp_right {a b} [decidable b] : a ∨ b ↔ (¬ b → a) :=
308+
or.comm.trans or_iff_not_imp_left
309+
310+
@[trans]
311+
lemma iff.trans {a b c} (h₁ : a ↔ b) (h₂ : b ↔ c) : a ↔ c :=
312+
iff.intro
313+
(assume ha, iff.mp h₂ (iff.mp h₁ ha))
314+
(assume hc, iff.mpr h₁ (iff.mpr h₂ hc))
315+
316+
theorem imp.swap {a b c : Prop} : (a → b → c) ↔ (b → a → c) :=
317+
⟨function.swap, function.swap⟩
318+
319+
theorem imp_not_comm {a b} : (a → ¬b) ↔ (b → ¬a) :=
320+
imp.swap
321+
322+
theorem not_and_of_not_or_not {a b} (h : ¬ a ∨ ¬ b) : ¬ (a ∧ b)
323+
| ⟨ha, hb⟩ := or.elim h (absurd ha) (absurd hb)
324+
325+
theorem not_and_distrib {a b} [decidable a] : ¬ (a ∧ b) ↔ ¬a ∨ ¬b :=
326+
⟨λ h, if ha : a then or.inr (λ hb, h ⟨ha, hb⟩) else or.inl ha, not_and_of_not_or_not⟩
327+
328+
theorem not_and_distrib' {a b} [decidable b] : ¬ (a ∧ b) ↔ ¬a ∨ ¬b :=
329+
⟨λ h, if hb : b then or.inl (λ ha, h ⟨ha, hb⟩) else or.inr hb, not_and_of_not_or_not⟩
330+
331+
theorem not_or_distrib {a b} : ¬ (a ∨ b) ↔ ¬ a ∧ ¬ b :=
332+
⟨λ h, ⟨λ ha, h (or.inl ha), λ hb, h (or.inr hb)⟩,
333+
λ ⟨h₁, h₂⟩ h, or.elim h h₁ h₂⟩
334+
335+
336+
meta def de_morgan_hyps : tactic unit :=
337+
do hs ← local_context,
338+
hs.for_each $ λ h,
339+
replace (some h.local_pp_name) none ``(not_and_distrib'.mp %%h) <|>
340+
replace (some h.local_pp_name) none ``(not_and_distrib.mp %%h) <|>
341+
replace (some h.local_pp_name) none ``(not_or_distrib.mp %%h) <|>
342+
skip
343+
304344
/--
305345
`tautology` breaks down assumptions of the form `_ ∧ _`, `_ ∨ _`, `_ ↔ _` and `∃ _, _`
306346
and splits a goal of the form `_ ∧ _`, `_ ↔ _` or `∃ _, _` until it can be discharged
@@ -310,12 +350,16 @@ meta def tautology : tactic unit :=
310350
repeat (do
311351
gs ← get_goals,
312352
() <$ tactic.intros;
353+
de_morgan_hyps,
313354
casesm (some ()) [``(_ ∧ _),``(_ ∨ _),``(Exists _),``(false)];
314355
constructor_matching (some ()) [``(_ ∧ _),``(_ ↔ _),``(true)],
356+
try (refine ``( or_iff_not_imp_left.mpr _)),
357+
try (refine ``( or_iff_not_imp_right.mpr _)),
315358
gs' ← get_goals,
316359
guard (gs ≠ gs') ) ;
317360
repeat
318-
(reflexivity <|> solve_by_elim <|> constructor_matching none [``(_ ∧ _),``(_ ↔ _),``(Exists _)]) ;
361+
(reflexivity <|> solve_by_elim <|>
362+
constructor_matching none [``(_ ∧ _),``(_ ↔ _),``(Exists _)] ) ;
319363
done
320364

321365
/-- Shorter name for the tactic `tautology`. -/

tests/examples.lean

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -97,4 +97,7 @@ end
9797

9898
example (p : Prop) : p ∧ true ↔ p := by tauto
9999

100-
example (p : Prop) : p ∨ false → p := by tauto
100+
example (p : Prop) : p ∨ false ↔ p := by tauto
101+
-- local attribute [instance] classical.prop_decidable
102+
example (p q r : Prop) [decidable p] [decidable r] : p ∨ (q ∧ r) ↔ (p ∨ q) ∧ (r ∨ p ∨ r) := by tauto
103+
example (p q r : Prop) [decidable q] [decidable r] : p ∨ (q ∧ r) ↔ (p ∨ q) ∧ (r ∨ p ∨ r) := by tauto

0 commit comments

Comments
 (0)