diff --git a/doc/changes.md b/doc/changes.md index 652a9b34f7..ce2b8f9d59 100644 --- a/doc/changes.md +++ b/doc/changes.md @@ -44,7 +44,13 @@ For more details, see discussion [here](https://github.com/leanprover/lean/pull/ * Add configuration object `rewrite_cfg` to `rw`/`rewrite` tactic. It now has support for `auto_param` and `opt_param`. The new goals are ordered using the same strategies available for `apply`. -* All `dsimp` tactics fail if they did not change anything. We can simulate the v3.2.0 by using `dsimp {fail_if_unchaged := ff}`. +* All `dsimp` tactics fail if they did not change anything. + We can simulate the v3.2.0 using `dsimp {fail_if_unchaged := ff}` or `try dsimp`. + +* All `dunfold` and `unfold` tactics fail if they did not unfold anything. + We can simulate the v3.2.0 using `unfold f {fail_if_unchaged := ff}` or `try {unfold f}`. + +* `dunfold_occs` was deleted, the new `conv` tactical should be used instead. v3.2.0 (18 June 2017) ------------- diff --git a/library/data/bitvec.lean b/library/data/bitvec.lean index ae7e8328f0..8f1d37243e 100644 --- a/library/data/bitvec.lean +++ b/library/data/bitvec.lean @@ -171,7 +171,7 @@ section conversion revert x, simp, induction xs with x xs ; intro y, { simp, unfold list.foldl add_lsb, simp [nat.mul_succ] }, - { simp, unfold list.foldl, apply ih_1 } + { simp, apply ih_1 } end theorem bits_to_nat_to_bool (n : ℕ) diff --git a/library/data/buffer/parser.lean b/library/data/buffer/parser.lean index ff811e657b..f7baa40777 100644 --- a/library/data/buffer/parser.lean +++ b/library/data/buffer/parser.lean @@ -40,8 +40,8 @@ begin apply funext, intro input, apply funext, intro pos, dunfold parser.bind, -cases (p input pos); dunfold bind, -cases (q result input pos_1); dunfold bind, +cases (p input pos); try {dunfold bind}, +cases (q result input pos_1); try {dunfold bind}, all_goals {refl} end diff --git a/library/init/data/int/bitwise.lean b/library/init/data/int/bitwise.lean index 2bca0f6bd3..c2d58270e4 100644 --- a/library/init/data/int/bitwise.lean +++ b/library/init/data/int/bitwise.lean @@ -183,7 +183,7 @@ namespace int all_goals { unfold cond, rw nat.bitwise_bit, repeat { rw bit_coe_nat <|> rw bit_neg_succ <|> rw bnot_bnot } }, - all_goals { unfold bnot; rw h; refl } + all_goals { unfold bnot {fail_if_unchaged := ff}; rw h; refl } end @[simp] lemma lor_bit (a m b n) : lor (bit a m) (bit b n) = bit (a || b) (lor m n) := diff --git a/library/init/meta/interactive.lean b/library/init/meta/interactive.lean index 2e8b0a10be..44164514a2 100644 --- a/library/init/meta/interactive.lean +++ b/library/init/meta/interactive.lean @@ -874,39 +874,28 @@ private meta def to_qualified_names : list name → tactic (list name) | [] := return [] | (c::cs) := do new_c ← to_qualified_name c, new_cs ← to_qualified_names cs, return (new_c::new_cs) -private meta def dunfold_hyps : list name → list name → tactic unit -| cs [] := skip -| cs (h::hs) := get_local h >>= dunfold_at cs >> dunfold_hyps cs hs +private meta def dunfold_hyps (cs : list name) (cfg : dunfold_config) : list name → tactic unit +| [] := skip +| (h::hs) := do h' ← get_local h, dunfold_at cs h' cfg, dunfold_hyps hs -meta def dunfold : parse ident* → parse location → tactic unit -| cs (loc.ns []) := do new_cs ← to_qualified_names cs, tactic.dunfold new_cs -| cs (loc.ns hs) := do new_cs ← to_qualified_names cs, dunfold_hyps new_cs hs -| cs (loc.wildcard) := do ls ← tactic.local_context, +meta def dunfold (cs : parse ident*) (l : parse location) (cfg : dunfold_config := {}) : tactic unit := +match l with +| (loc.ns []) := do new_cs ← to_qualified_names cs, tactic.dunfold new_cs cfg +| (loc.ns hs) := do new_cs ← to_qualified_names cs, dunfold_hyps new_cs cfg hs +| (loc.wildcard) := do ls ← tactic.local_context, n ← revert_lst ls, new_cs ← to_qualified_names cs, - tactic.dunfold new_cs, + tactic.dunfold new_cs cfg, intron n +end /- TODO(Leo): add support for non-refl lemmas -/ -meta def unfold : parse ident* → parse location → tactic unit := -dunfold - -private meta def dunfold_hyps_occs : name → occurrences → list name → tactic unit -| c occs [] := skip -| c occs (h::hs) := get_local h >>= dunfold_core_at occs [c] >> dunfold_hyps_occs c occs hs - -meta def dunfold_occs : parse ident → parse location → list nat → tactic unit -| c (loc.ns []) ps := do new_c ← to_qualified_name c, tactic.dunfold_occs_of ps new_c -| c (loc.ns hs) ps := do new_c ← to_qualified_name c, dunfold_hyps_occs new_c (occurrences.pos ps) hs -| c (loc.wildcard) ps := fail "wildcard not allowed when unfolding occurences" - -/- TODO(Leo): add support for non-refl lemmas -/ -meta def unfold_occs : parse ident → parse location → list nat → tactic unit := -dunfold_occs +meta def unfold (cs : parse ident*) (l : parse location) (cfg : dunfold_config := {}) : tactic unit := +dunfold cs l cfg private meta def delta_hyps : list name → list name → tactic unit | cs [] := skip -| cs (h::hs) := get_local h >>= delta_at cs >> dunfold_hyps cs hs +| cs (h::hs) := get_local h >>= delta_at cs >> delta_hyps cs hs meta def delta : parse ident* → parse location → tactic unit | cs (loc.ns []) := do new_cs ← to_qualified_names cs, tactic.delta new_cs diff --git a/library/init/meta/simp_tactic.lean b/library/init/meta/simp_tactic.lean index 8626516b19..61c04b236b 100644 --- a/library/init/meta/simp_tactic.lean +++ b/library/init/meta/simp_tactic.lean @@ -107,30 +107,22 @@ meta def dsimplify (λ u e, do r ← post e, return (u, r)) e, return new_e -meta constant dunfold_expr_core : transparency → expr → tactic expr - /- Remark: we use transparency.instances by default to make sure that we can unfold projections of type classes. Example: - dunfold_expr (@has_add.add nat nat.has_add a b) + (@has_add.add nat nat.has_add a b) -/ -meta def dunfold_expr : expr → tactic expr := -dunfold_expr_core transparency.instances - -meta constant unfold_projection_core : transparency → expr → tactic expr +/-- If `e` is a projection application, try to unfold it, otherwise fail. -/ +meta constant unfold_projection (e : expr) (md : transparency := transparency.instances) : tactic expr -meta def unfold_projection : expr → tactic expr := -unfold_projection_core transparency.instances +structure dunfold_config extends dsimp_config := +(md := transparency.instances) -meta constant dunfold_occs_core (m : transparency) (max_steps : nat) (occs : occurrences) (cs : list name) (e : expr) : tactic expr -meta constant dunfold_core (m : transparency) (max_steps : nat) (cs : list name) (e : expr) : tactic expr +meta constant dunfold_core (cs : list name) (e : expr) (cfg : dunfold_config := {}) : tactic expr -meta def dunfold : list name → tactic unit := -λ cs, target >>= dunfold_core transparency.instances default_max_steps cs >>= unsafe_change - -meta def dunfold_occs_of (occs : list nat) (c : name) : tactic unit := -target >>= dunfold_occs_core transparency.instances default_max_steps (occurrences.pos occs) [c] >>= unsafe_change +meta def dunfold (cs : list name) (cfg : dunfold_config := {}) : tactic unit := +do t ← target, dunfold_core cs t cfg >>= unsafe_change meta def revert_and_transform (transform : expr → tactic expr) (h : expr) : tactic unit := do num_reverted : ℕ ← revert h, @@ -146,11 +138,8 @@ do num_reverted : ℕ ← revert h, end, intron num_reverted -meta def dunfold_core_at (occs : occurrences) (cs : list name) : expr → tactic unit := -revert_and_transform (dunfold_occs_core transparency.instances default_max_steps occs cs) - -meta def dunfold_at (cs : list name) : expr → tactic unit := -revert_and_transform (dunfold_core transparency.instances default_max_steps cs) +meta def dunfold_at (cs : list name) (h : expr) (cfg : dunfold_config := {}) : tactic unit := +revert_and_transform (λ e, dunfold_core cs e cfg) h structure delta_config := (max_steps := default_max_steps) @@ -164,7 +153,7 @@ cs.any (λ c, f.is_constant && f.const_name.is_internal && (f.const_name.get_prefix = c)) /-- Delta reduce the given constant names -/ -meta def delta_core (cfg : delta_config) (cs : list name) (e : expr) : tactic expr := +meta def delta_expr (cs : list name) (e : expr) (cfg : delta_config) : tactic expr := let unfold (u : unit) (e : expr) : tactic (unit × expr × bool) := do guard (is_delta_target e cs), (expr.const f_name f_lvls) ← return e.get_app_fn, @@ -177,23 +166,23 @@ in do (c, new_e) ← dsimplify_core () (λ c e, failed) unfold e {max_steps := c return new_e meta def delta (cs : list name) : tactic unit := -target >>= delta_core {} cs >>= unsafe_change +do t ← target, delta_expr cs t {} >>= unsafe_change meta def delta_at (cs : list name) : expr → tactic unit := -revert_and_transform (delta_core {} cs) +revert_and_transform (λ e, delta_expr cs e {}) -meta def unfold_projections_core (m : transparency) (max_steps : nat) (e : expr) : tactic expr := +meta def unfold_projections_core (e : expr) (md := transparency.instances) (max_steps : nat := default_max_steps) : tactic expr := let unfold (changed : bool) (e : expr) : tactic (bool × expr × bool) := do - new_e ← unfold_projection_core m e, + new_e ← unfold_projection e md, return (tt, new_e, tt) -in do (tt, new_e) ← dsimplify_core ff (λ c e, failed) unfold e | fail "no projections to unfold", +in do (tt, new_e) ← dsimplify_core ff (λ c e, failed) unfold e {max_steps := max_steps, md := md} | fail "no projections to unfold", return new_e -meta def unfold_projections : tactic unit := -target >>= unfold_projections_core semireducible default_max_steps >>= change +meta def unfold_projections (md := transparency.instances) (max_steps : nat := default_max_steps) : tactic unit := +do t ← target, unfold_projections_core t md max_steps >>= unsafe_change -meta def unfold_projections_at : expr → tactic unit := -revert_and_transform (unfold_projections_core semireducible default_max_steps) +meta def unfold_projections_at (h : expr) (md := transparency.instances) (max_steps : nat := default_max_steps) : tactic unit := +revert_and_transform (λ e, unfold_projections_core e md max_steps) h structure simp_config := (max_steps : nat := default_max_steps) diff --git a/library/init/meta/well_founded_tactics.lean b/library/init/meta/well_founded_tactics.lean index 92ad74b9af..565e7965b7 100644 --- a/library/init/meta/well_founded_tactics.lean +++ b/library/init/meta/well_founded_tactics.lean @@ -36,7 +36,7 @@ meta def clear_internals : tactic unit := local_context >>= clear_wf_rec_goal_aux meta def unfold_wf_rel : tactic unit := -dunfold [``has_well_founded.r] +dunfold [``has_well_founded.r] {fail_if_unchaged := ff} meta def is_psigma_mk : expr → tactic (expr × expr) | `(psigma.mk %%a %%b) := return (a, b) @@ -57,7 +57,7 @@ meta def process_lex : tactic unit → tactic unit tac private meta def unfold_sizeof_measure : tactic unit := -dunfold [``sizeof_measure, ``measure, ``inv_image] +dunfold [``sizeof_measure, ``measure, ``inv_image] {fail_if_unchaged := ff} private meta def add_simps : simp_lemmas → list name → tactic simp_lemmas | s [] := return s @@ -77,7 +77,7 @@ e.mfold simp_lemmas.mk $ λ c d s, private meta def unfold_sizeof_loop : tactic unit := do - dunfold [``sizeof, ``has_sizeof.sizeof], + dunfold [``sizeof, ``has_sizeof.sizeof] {fail_if_unchaged := ff}, S ← target >>= collect_sizeof_lemmas, (simplify_goal S >> unfold_sizeof_loop) <|> diff --git a/src/library/tactic/unfold_tactic.cpp b/src/library/tactic/unfold_tactic.cpp index 5b4df64af1..a90eb81082 100644 --- a/src/library/tactic/unfold_tactic.cpp +++ b/src/library/tactic/unfold_tactic.cpp @@ -18,7 +18,7 @@ Author: Leonardo de Moura #include "library/tactic/dsimplify.h" namespace lean { -vm_obj tactic_unfold_projection_core(vm_obj const & m, vm_obj const & _e, vm_obj const & _s) { +vm_obj tactic_unfold_projection(vm_obj const & _e, vm_obj const & m, vm_obj const & _s) { expr const & e = to_expr(_e); tactic_state const & s = tactic::to_state(_s); try { @@ -60,42 +60,6 @@ optional dunfold(type_context & ctx, expr const & e) { } } -vm_obj tactic_dunfold_expr_core(vm_obj const & m, vm_obj const & _e, vm_obj const & _s) { - expr const & e = to_expr(_e); - tactic_state const & s = tactic::to_state(_s); - try { - environment const & env = s.env(); - expr const & fn = get_app_fn(e); - if (!is_constant(fn)) - return tactic::mk_exception("dunfold_expr failed, expression is not a constant nor a constant application", s); - if (is_projection(s.env(), const_name(fn))) { - type_context ctx = mk_type_context_for(s, to_transparency_mode(m)); - if (auto new_e = ctx.reduce_projection(e)) - return tactic::mk_success(to_obj(*new_e), s); - return tactic::mk_exception("dunfold_expr failed, failed to unfold projection", s); - } else if (has_eqn_lemmas(env, const_name(fn))) { - type_context ctx = mk_type_context_for(s, to_transparency_mode(m)); - if (auto new_e = dunfold(ctx, e)) { - return tactic::mk_success(to_obj(*new_e), s); - } else { - return tactic::mk_exception("dunfold_expr failed, none of the rfl lemmas is applicable", s); - } - } else if (auto new_e = unfold_term(env, e)) { - return tactic::mk_success(to_obj(*new_e), s); - } else { - return tactic::mk_exception("dunfold_expr failed, failed to unfold", s); - } - } catch (exception & ex) { - return tactic::mk_exception(ex, s); - } -} - -static dsimp_config mk_dsimp_cfg(unsigned max_steps) { - dsimp_config cfg; - cfg.m_max_steps = max_steps; - return cfg; -} - class unfold_core_fn : public dsimplify_core_fn { protected: name_set m_cs; @@ -125,12 +89,12 @@ class unfold_core_fn : public dsimplify_core_fn { return none(); if (!check_occ()) return none(); - return optional>(*new_e, true); + return optional>(*new_e, !m_cfg.m_single_pass); } public: - unfold_core_fn(type_context & ctx, defeq_canonizer::state & dcs, unsigned max_steps, list const & cs): - dsimplify_core_fn(ctx, dcs, mk_dsimp_cfg(max_steps)), + unfold_core_fn(type_context & ctx, defeq_canonizer::state & dcs, list const & cs, dsimp_config const & cfg): + dsimplify_core_fn(ctx, dcs, cfg), m_cs(to_name_set(cs)) { } }; @@ -140,69 +104,38 @@ class unfold_fn : public unfold_core_fn { return unfold_step(e); } public: - unfold_fn(type_context & ctx, defeq_canonizer::state & dcs, unsigned max_steps, list const & cs): - unfold_core_fn(ctx, dcs, max_steps, cs) { - } -}; - -class unfold_occs_fn : public unfold_core_fn { - occurrences m_occs; - unsigned m_counter{1}; - - virtual bool check_occ() override { - bool r = m_occs.contains(m_counter); - m_counter++; - return r; - } - - virtual optional> pre(expr const & e) override { - return unfold_step(e); - } - -public: - unfold_occs_fn(type_context & ctx, defeq_canonizer::state & dcs, unsigned max_steps, - occurrences const & occs, list const & cs): - unfold_core_fn(ctx, dcs, max_steps, cs), - m_occs(occs) { + unfold_fn(type_context & ctx, defeq_canonizer::state & dcs, list const & cs, dsimp_config const & cfg): + unfold_core_fn(ctx, dcs, cs, cfg) { } }; -vm_obj tactic_dunfold_core(vm_obj const & m, vm_obj const & max_steps, vm_obj const & cs, vm_obj const & _e, vm_obj const & _s) { +vm_obj tactic_dunfold_core(vm_obj const & cs, vm_obj const & _e, vm_obj const & _cfg, vm_obj const & _s) { expr const & e = to_expr(_e); tactic_state const & s = tactic::to_state(_s); defeq_can_state dcs = s.dcs(); - type_context ctx = mk_type_context_for(s, to_transparency_mode(m)); - unfold_fn F(ctx, dcs, force_to_unsigned(max_steps), to_list_name(cs)); + /* + structure dunfold_config extends dsimp_config := + (md := transparency.instances) -- this is not a new field. + */ + dsimp_config cfg(_cfg); + type_context ctx = mk_type_context_for(s, cfg.m_md); + unfold_fn F(ctx, dcs, to_list_name(cs), cfg); try { expr new_e = F(e); tactic_state new_s = set_mctx_dcs(s, F.mctx(), dcs); - return tactic::mk_success(to_obj(new_e), new_s); - } catch (exception & ex) { - return tactic::mk_exception(ex, s); - } -} - -vm_obj tactic_dunfold_occs_core(vm_obj const & m, vm_obj const & max_steps, vm_obj const & occs, vm_obj const & cs, - vm_obj const & _e, vm_obj const & _s) { - expr const & e = to_expr(_e); - tactic_state const & s = tactic::to_state(_s); - defeq_can_state dcs = s.dcs(); - type_context ctx = mk_type_context_for(s, to_transparency_mode(m)); - unfold_occs_fn F(ctx, dcs, force_to_unsigned(max_steps), to_occurrences(occs), to_list_name(cs)); - try { - expr new_e = F(e); - tactic_state new_s = set_mctx_dcs(s, F.mctx(), dcs); - return tactic::mk_success(to_obj(new_e), new_s); + if (cfg.m_fail_if_unchanged && e == new_e) { + return tactic::mk_exception("dunfold tactic failed to unfold", s); + } else { + return tactic::mk_success(to_obj(new_e), new_s); + } } catch (exception & ex) { return tactic::mk_exception(ex, s); } } void initialize_unfold_tactic() { - DECLARE_VM_BUILTIN(name({"tactic", "unfold_projection_core"}), tactic_unfold_projection_core); - DECLARE_VM_BUILTIN(name({"tactic", "dunfold_expr_core"}), tactic_dunfold_expr_core); - DECLARE_VM_BUILTIN(name({"tactic", "dunfold_core"}), tactic_dunfold_core); - DECLARE_VM_BUILTIN(name({"tactic", "dunfold_occs_core"}), tactic_dunfold_occs_core); + DECLARE_VM_BUILTIN(name({"tactic", "unfold_projection"}), tactic_unfold_projection); + DECLARE_VM_BUILTIN(name({"tactic", "dunfold_core"}), tactic_dunfold_core); } void finalize_unfold_tactic() { diff --git a/tests/lean/run/dunfold1.lean b/tests/lean/run/dunfold1.lean deleted file mode 100644 index 9a60f44ed2..0000000000 --- a/tests/lean/run/dunfold1.lean +++ /dev/null @@ -1,23 +0,0 @@ -open tactic - -set_option pp.all true - -def g : nat → nat := λ x, x + 5 - -example (a b : nat) (p : nat → Prop) (h : p (g (nat.succ (nat.succ a)))) : p (g (a + 2)) := -by do - t ← target, - new_t ← dsimplify (λ e, failed) - (λ e, - do { new_e ← unfold_projection e, return (new_e, tt) } - <|> - do { - guard ([`add, `nat.add, `one, `zero, `bit0, `bit1]^.any e^.is_app_of), - new_e ← dunfold_expr e, trace e, - trace "===>", trace new_e, trace "-------", - return (new_e, tt) }) - t, - expected ← to_expr ```(p (g (nat.succ (nat.succ a)))), - guard (new_t = expected), - trace new_t, - assumption diff --git a/tests/lean/run/dunfold2.lean b/tests/lean/run/dunfold2.lean deleted file mode 100644 index 73aacd722c..0000000000 --- a/tests/lean/run/dunfold2.lean +++ /dev/null @@ -1,52 +0,0 @@ -open tactic - -set_option pp.all true - -def g : nat → nat := λ x, x + 5 - -example (a b : nat) (p : nat → Prop) (h : p (g (nat.succ (nat.succ a)))) : p (g (a + 2)) := -by do - t ← target, - new_t ← dsimplify (λ e, failed) - (λ e, - do { new_e ← unfold_projection e, return (new_e, tt) } - <|> - do { - guard ([`add, `nat.add, `one, `zero]^.any e^.is_app_of), - /- We are using transparency.none, so - - nat.add a (bit0 (nat.succ nat.zero)) - - cannot be matched with - - nat.add ?a (nat.succ ?b) - -/ - new_e ← dunfold_expr_core transparency.none e, trace e, - trace "===>", trace new_e, trace "-------", - return (new_e, tt) }) - t, - trace new_t, - expected ← to_expr ```(p (g (nat.add a (bit0 (nat.succ nat.zero))))), - guard (new_t = expected), - trace new_t, - assumption - - -example (a b : nat) (p : nat → Prop) (h : p (g (nat.succ (nat.succ a)))) : p (g (a + nat.succ (nat.succ 0))) := -by do - t ← target, - new_t ← dsimplify (λ e, failed) - (λ e, - do { new_e ← unfold_projection e, return (new_e, tt) } - <|> - do { - guard ([`add, `nat.add, `one, `zero]^.any e^.is_app_of), - new_e ← dunfold_expr_core transparency.none e, trace e, - trace "===>", trace new_e, trace "-------", - return (new_e, tt) }) - t, - trace new_t, - expected ← to_expr ```(p (g (nat.succ (nat.succ a)))), - guard (new_t = expected), - trace new_t, - assumption diff --git a/tests/lean/run/dunfold3.lean b/tests/lean/run/dunfold3.lean index a5d7e45e15..6febb67510 100644 --- a/tests/lean/run/dunfold3.lean +++ b/tests/lean/run/dunfold3.lean @@ -14,17 +14,3 @@ end meta def check_expected (p : pexpr) : tactic unit := do t ← target, ex ← to_expr p, guard (t = ex) - -example (a b c : nat) (f : nat → nat → nat) (h : false) : f (g a) (g b) = (g c) := -begin - unfold_occs g [2], - check_expected ```(f (g a) (b + 5) = g c), - contradiction -end - -example (a b c : nat) (f : nat → nat → nat) (h : false) : f (g a) (g b) = (g c) := -begin - unfold_occs g [1, 3], - check_expected ```(f (a + 5) (g b) = c + 5), - contradiction -end diff --git a/tests/lean/run/unfold_crash.lean b/tests/lean/run/unfold_crash.lean deleted file mode 100644 index a5b3ab0f28..0000000000 --- a/tests/lean/run/unfold_crash.lean +++ /dev/null @@ -1,40 +0,0 @@ -open tactic - --- TODO(Leo): remove the following command. --- It has been added to avoid a crash in the code generator -noncomputable theory - -inductive vector (A : Type) : nat → Type - | nil {} : vector 0 - | cons : Π {n}, A -> vector n -> vector (nat.succ n) - -definition vmap {A B : Type} (f : A -> B) : Π {n}, vector A n -> vector B n -| 0 vector.nil := vector.nil -| (n+1) (vector.cons x xs) := vector.cons (f x) (vmap xs) - -noncomputable definition vappend {A} : Π {n m}, vector A n → vector A m → vector A (m + n) -| 0 0 vector.nil vector.nil := vector.nil -| 0 (m+1) vector.nil (vector.cons x xs) := vector.cons x xs -| (n+1) 0 (vector.cons x xs) vector.nil := vector.cons x (vappend xs vector.nil) -| (n+1) (m+1) (vector.cons x xs) (vector.cons y ys) := vector.cons x (vappend xs (vector.cons y ys)) - -axiom Sorry : ∀ A, A - -theorem vappend_assoc : - Π {A : Type} {n m k : nat} (v1 : vector A n) (v2 : vector A m) (v3 : vector A k), - vappend (vappend v1 v2) v3 == vappend v1 (vappend v2 v3) := -by do - intros, - v <- get_local `v1, - induction v, - v2 ← get_local `v2, - cases v2 [`m, `h2, `t2], - trace_state, trace "------", - -- unfold only the first occurrence (i.e., the one of the form (vappend nil nil) - dunfold_occs_of [1] `vappend, - trace_state, trace "------", - mk_const `Sorry >>= apply, - -- unfold only the first occurrence (i.e., the one of the form (vappend nil (cons ...)) - dunfold_occs_of [1] `vappend, - trace_state, trace "------", - repeat $ mk_const `Sorry >>= apply