Skip to content
Snippets Groups Projects
Commit d9536025 authored by Ralf Jung's avatar Ralf Jung
Browse files

add FIXMEs for Coq bug work-arounds

parent abbd8d0e
No related branches found
No related tags found
No related merge requests found
...@@ -104,6 +104,7 @@ Proof. ...@@ -104,6 +104,7 @@ Proof.
intros (l&?&?). exists (hash x, l); simpl. by rewrite elem_of_map_to_list. intros (l&?&?). exists (hash x, l); simpl. by rewrite elem_of_map_to_list.
- unfold elements, hashset_elements. intros [m Hm]; simpl. - unfold elements, hashset_elements. intros [m Hm]; simpl.
rewrite map_Forall_to_list in Hm. generalize (NoDup_fst_map_to_list m). rewrite map_Forall_to_list in Hm. generalize (NoDup_fst_map_to_list m).
(* FIXME: trailing [?] works around Coq bug #12944. *)
induction Hm as [|[n l] m' [??] Hm ?]; induction Hm as [|[n l] m' [??] Hm ?];
csimpl; inversion_clear 1 as [|?? Hn]; [constructor|]. csimpl; inversion_clear 1 as [|?? Hn]; [constructor|].
apply NoDup_app; split_and?; eauto. apply NoDup_app; split_and?; eauto.
......
...@@ -180,6 +180,7 @@ Section merge_sort_correct. ...@@ -180,6 +180,7 @@ Section merge_sort_correct.
Sorted R l1 Sorted R l2 Sorted R (list_merge R l1 l2). Sorted R l1 Sorted R l2 Sorted R (list_merge R l1 l2).
Proof. Proof.
intros Hl1. revert l2. induction Hl1 as [|x1 l1 IH1]; intros Hl1. revert l2. induction Hl1 as [|x1 l1 IH1];
(* FIXME: trailing [?] works around Coq bug #12944. *)
induction 1 as [|x2 l2 IH2 ?]; rewrite ?list_merge_cons; simpl; induction 1 as [|x2 l2 IH2 ?]; rewrite ?list_merge_cons; simpl;
repeat case_decide; repeat case_decide;
repeat match goal with H : ¬R _ _ |- _ => apply total_not in H end; repeat match goal with H : ¬R _ _ |- _ => apply total_not in H end;
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment