blob: db4b69dde40b90e57519cff53cce003111ca7c58 (
plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
|
Require Import Coq.Classes.Morphisms.
Global Instance option_rect_Proper_nd {A T}
: Proper ((pointwise_relation _ eq) ==> eq ==> eq ==> eq) (@option_rect A (fun _ => T)).
Proof.
intros ?? H ??? [|]??; subst; simpl; congruence.
Qed.
Global Instance option_rect_Proper_nd' {A T}
: Proper ((pointwise_relation _ eq) ==> eq ==> forall_relation (fun _ => eq)) (@option_rect A (fun _ => T)).
Proof.
intros ?? H ??? [|]; subst; simpl; congruence.
Qed.
Hint Extern 1 (Proper _ (@option_rect ?A (fun _ => ?T))) => exact (@option_rect_Proper_nd' A T) : typeclass_instances.
Lemma option_rect_option_map : forall {A B C} (f:A->B) some none v,
option_rect (fun _ => C) (fun x => some (f x)) none v = option_rect (fun _ => C) some none (option_map f v).
Proof.
destruct v; reflexivity.
Qed.
Lemma option_rect_function {A B C S' N' v} f
: f (option_rect (fun _ : option A => option B) S' N' v)
= option_rect (fun _ : option A => C) (fun x => f (S' x)) (f N') v.
Proof. destruct v; reflexivity. Qed.
(*
Ltac commute_option_rect_Let_In := (* pull let binders out side of option_rect pattern matching *)
idtac;
lazymatch goal with
| [ |- ?LHS = option_rect ?P ?S ?N (Let_In ?x ?f) ]
=> (* we want to just do a [change] here, but unification is stupid, so we have to tell it what to unfold in what order *)
cut (LHS = Let_In x (fun y => option_rect P S N (f y))); cbv beta;
[ set_evars;
let H := fresh in
intro H;
rewrite H;
clear;
abstract (cbv [Let_In]; reflexivity)
| ]
end.
*)
(** TODO: possibly move me, remove local *)
Ltac replace_option_match_with_option_rect :=
idtac;
lazymatch goal with
| [ |- _ = ?RHS :> ?T ]
=> lazymatch RHS with
| match ?a with None => ?N | Some x => @?S x end
=> replace RHS with (option_rect (fun _ => T) S N a) by (destruct a; reflexivity)
end
end.
Ltac simpl_option_rect := (* deal with [option_rect _ _ _ None] and [option_rect _ _ _ (Some _)] *)
repeat match goal with
| [ |- context[option_rect ?P ?S ?N None] ]
=> change (option_rect P S N None) with N
| [ |- context[option_rect ?P ?S ?N (Some ?x) ] ]
=> change (option_rect P S N (Some x)) with (S x); cbv beta
end.
|