feat(library,hott): add notation T1 : T2
as syntax sugar for (focus (T1; all_goals T2))
closes #775
This commit is contained in:
parent
938dae7b19
commit
582dbecfd0
5 changed files with 42 additions and 1 deletions
|
@ -148,4 +148,5 @@ nat.rec id (λn t', and_then t t') (nat.of_num n)
|
||||||
|
|
||||||
end tactic
|
end tactic
|
||||||
tactic_infixl `;`:15 := tactic.and_then
|
tactic_infixl `;`:15 := tactic.and_then
|
||||||
|
tactic_notation T1 `:`:15 T2 := tactic.focus (tactic.and_then T1 (tactic.all_goals T2))
|
||||||
tactic_notation `(` h `|` r:(foldl `|` (e r, tactic.or_else r e) h) `)` := r
|
tactic_notation `(` h `|` r:(foldl `|` (e r, tactic.or_else r e) h) `)` := r
|
||||||
|
|
|
@ -147,4 +147,5 @@ definition do (n : num) (t : tactic) : tactic :=
|
||||||
nat.rec id (λn t', and_then t t') (nat.of_num n)
|
nat.rec id (λn t', and_then t t') (nat.of_num n)
|
||||||
end tactic
|
end tactic
|
||||||
tactic_infixl `;`:15 := tactic.and_then
|
tactic_infixl `;`:15 := tactic.and_then
|
||||||
|
tactic_notation T1 `:`:15 T2 := tactic.focus (tactic.and_then T1 (tactic.all_goals T2))
|
||||||
tactic_notation `(` h `|` r:(foldl `|` (e r, tactic.or_else r e) h) `)` := r
|
tactic_notation `(` h `|` r:(foldl `|` (e r, tactic.or_else r e) h) `)` := r
|
||||||
|
|
|
@ -3,7 +3,7 @@ definition id_1 (n : nat) :=
|
||||||
by exact n
|
by exact n
|
||||||
|
|
||||||
definition id_2 (n : nat) :=
|
definition id_2 (n : nat) :=
|
||||||
(by exact n : nat)
|
((by exact n) : nat)
|
||||||
|
|
||||||
definition id_3 (n : nat) : nat :=
|
definition id_3 (n : nat) : nat :=
|
||||||
by exact n
|
by exact n
|
||||||
|
|
10
tests/lean/775.lean
Normal file
10
tests/lean/775.lean
Normal file
|
@ -0,0 +1,10 @@
|
||||||
|
open nat
|
||||||
|
|
||||||
|
tactic_notation T1 `:`:15 T2 := tactic.focus (tactic.and_then T1 (tactic.all_goals T2))
|
||||||
|
|
||||||
|
example (P Q : ℕ → ℕ → Prop) (n m : ℕ) (p : P n m) : Q n m ∧ false :=
|
||||||
|
begin
|
||||||
|
split,
|
||||||
|
revert m p, induction n : intro m; induction m : intro p,
|
||||||
|
state, repeat exact sorry
|
||||||
|
end
|
29
tests/lean/775.lean.expected.out
Normal file
29
tests/lean/775.lean.expected.out
Normal file
|
@ -0,0 +1,29 @@
|
||||||
|
775.lean:9:2: proof state
|
||||||
|
P Q : ℕ → ℕ → Prop,
|
||||||
|
p : P 0 0
|
||||||
|
⊢ Q 0 0
|
||||||
|
|
||||||
|
P Q : ℕ → ℕ → Prop,
|
||||||
|
a : ℕ,
|
||||||
|
v_0 : P 0 a → Q 0 a,
|
||||||
|
p : P 0 (succ a)
|
||||||
|
⊢ Q 0 (succ a)
|
||||||
|
|
||||||
|
P Q : ℕ → ℕ → Prop,
|
||||||
|
a : ℕ,
|
||||||
|
v_0 : ∀ (m : ℕ), P a m → Q a m,
|
||||||
|
p : P (succ a) 0
|
||||||
|
⊢ Q (succ a) 0
|
||||||
|
|
||||||
|
P Q : ℕ → ℕ → Prop,
|
||||||
|
a : ℕ,
|
||||||
|
v_0 : ∀ (m : ℕ), P a m → Q a m,
|
||||||
|
a_1 : ℕ,
|
||||||
|
v_0_1 : P (succ a) a_1 → Q (succ a) a_1,
|
||||||
|
p : P (succ a) (succ a_1)
|
||||||
|
⊢ Q (succ a) (succ a_1)
|
||||||
|
|
||||||
|
P Q : ℕ → ℕ → Prop,
|
||||||
|
n m : ℕ,
|
||||||
|
p : P n m
|
||||||
|
⊢ false
|
Loading…
Reference in a new issue