Commit 8bb63778 authored by Adam Chlipala's avatar Adam Chlipala

Get Coinductive compiling again

parent e51986f6
......@@ -412,7 +412,7 @@ Definition regmapCompat (rm : regmap nat) (rm' : regmap (option nat)) :=
(** remove printing * *)
Ltac compat := unfold regmapCompat in *; crush;
match goal with
| [ |- match get _ ?R with Some _ => _ | None => _ end ] => destruct R; crush
| [ H : _ |- match get _ ?R with Some _ => _ | None => _ end ] => generalize (H R); destruct R; crush
end.
Lemma regmapCompat_set_None : forall rm rm' r n,
......
......@@ -111,17 +111,16 @@ Ltac un_done :=
| [ H : done _ |- _ ] => clear H
end.
Ltac crush'' tryLemmas lemmas invOne :=
Ltac crush' lemmas invOne :=
let sintuition := simpl in *; intuition; try subst; repeat (simplHyp invOne; intuition; try subst); try congruence
in (sintuition; rewriter;
match tryLemmas with
| true => repeat ((app ltac:(fun L => inster L L) lemmas || appHyps ltac:(fun L => inster L L));
match lemmas with
| false => idtac
| _ => repeat ((app ltac:(fun L => inster L L) lemmas || appHyps ltac:(fun L => inster L L));
repeat (simplHyp invOne; intuition)); un_done
| _ => idtac
end; sintuition; try omega; try (elimtype False; omega)).
Ltac crush := crush'' false tt fail.
Ltac crush' := crush'' true.
Ltac crush := crush' false fail.
Theorem dep_destruct : forall (T : Type) (T' : T -> Type) x (v : T' x) (P : T' x -> Prop),
(forall x' (v' : T' x') (Heq : x' = x), P (match Heq in (_ = x) return (T' x) with
......
Markdown is supported
0% or
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment