Skip to content
Snippets Groups Projects
Commit c9af4df5 authored by Janno's avatar Janno
Browse files

more updates; added infrastructure file for typeclasses

parent 2126ec9d
No related branches found
No related tags found
No related merge requests found
-Q . ra -Q . ra
infrastructure.v
machine.v machine.v
lang.v lang.v
tactics.v tactics.v
......
From mathcomp Require Import ssreflect ssrbool seq.
From iris.prelude Require Export strings list numbers.
From iris.prelude Require Export gmap finite mapset.
Global Generalizable All Variables.
Global Set Automatic Coercions Import.
Global Set Asymmetric Patterns.
Global Set Bullet Behavior "Strict Subproofs".
From Coq Require Export Utf8.
Instance injective_dec_eq
`{H : x y : A, Decision (x = y)} {B : Type}
f (g : A -> option B) (Inj : x, g (f x) = Some x) : x y : B, Decision (x = y) | 1000.
Proof.
move => x y. case: (decide (f x = f y)).
- move/(f_equal g). rewrite !Inj => [[]]. by left.
- move => NEq. right => Eq. apply: NEq. by rewrite Eq.
Qed.
Class Extension A := extension : relation A.
Instance: Params (@extension) 2.
Infix "⊑" := extension (at level 70) : C_scope.
Notation "(⊑)" := extension (only parsing) : C_scope.
Notation "( X ⊑ )" := (extension X) (only parsing) : C_scope.
Notation "( ⊑ X )" := (λ Y, Y X) (only parsing) : C_scope.
Infix "⊑*" := (Forall2 ()) (at level 70) : C_scope.
Notation "(⊑*)" := (Forall2 ()) (only parsing) : C_scope.
Infix "⊑**" := (Forall2 (*)) (at level 70) : C_scope.
Infix "⊑1*" := (Forall2 (λ p q, p.1 q.1)) (at level 70) : C_scope.
Infix "⊑2*" := (Forall2 (λ p q, p.2 q.2)) (at level 70) : C_scope.
Infix "⊑1**" := (Forall2 (λ p q, p.1 * q.1)) (at level 70) : C_scope.
Infix "⊑2**" := (Forall2 (λ p q, p.2 * q.2)) (at level 70) : C_scope.
Class StrictExtension A := strict_extension : relation A.
Instance: Params (@strict_extension) 2.
Infix "⊏" := strict_extension (at level 70) : C_scope.
Notation "(⊏)" := strict_extension (only parsing) : C_scope.
Notation "( X ⊏ )" := (strict_extension X) (only parsing) : C_scope.
Notation "( ⊏ X )" := (λ Y, Y X) (only parsing) : C_scope.
Infix "⊏*" := (Forall2 ()) (at level 70) : C_scope.
Notation "(⊏*)" := (Forall2 ()) (only parsing) : C_scope.
Infix "⊏**" := (Forall2 (*)) (at level 70) : C_scope.
Infix "⊏1*" := (Forall2 (λ p q, p.1 q.1)) (at level 70) : C_scope.
Infix "⊏2*" := (Forall2 (λ p q, p.2 q.2)) (at level 70) : C_scope.
Infix "⊏1**" := (Forall2 (λ p q, p.1 * q.1)) (at level 70) : C_scope.
Infix "⊏2**" := (Forall2 (λ p q, p.2 * q.2)) (at level 70) : C_scope.
Class Join A := join : A A A.
Instance: Params (@join) 2.
Infix "⊔" := join (at level 50, left associativity) : C_scope.
Notation "(⊔)" := join (only parsing) : C_scope.
Notation "( x ⊔)" := (join x) (only parsing) : C_scope.
Notation "(⊔ x )" := (λ y, join y x) (only parsing) : C_scope.
Infix "⊔*" := (zip_with ()) (at level 50, left associativity) : C_scope.
Notation "(⊔*)" := (zip_with ()) (only parsing) : C_scope.
Infix "⊔**" := (zip_with (zip_with ()))
(at level 50, left associativity) : C_scope.
Infix "⊔*⊔**" := (zip_with (prod_zip () (*)))
(at level 50, left associativity) : C_scope.
Class JoinSemiLattice A `{Extension A, Join A} : Prop := {
join_semi_lattice_pre :>> PreOrder ();
join_ext_l X Y : X X Y;
join_ext_r X Y : Y X Y;
join_least X Y Z : X Z Y Z X Y Z
}.
Instance JSL_Comm `{E : Extension A} `{J : Join A} `{@PartialOrder A ()} `{@JoinSemiLattice A E J} : @Comm A _ (=) ().
Proof.
move => ? ?.
apply: anti_symm; apply: join_least; by eauto using join_ext_l, join_ext_r.
Qed.
Instance JSL_Assoc `{E : Extension A} `{J : Join A} `{@PartialOrder A ()} `{@JoinSemiLattice A E J} : @Assoc A (=) ().
Proof.
move => ? ? ?.
apply: anti_symm; repeat apply: join_least;
by eauto using PreOrder_Transitive, join_ext_l, join_ext_r.
Qed.
(* Generic Instances for Extension, Join and JoinSemiLattice *)
Instance prod_Extension `{Extension A} `{Extension B} : Extension (A * B) :=
λ p1 p2, p1.1 p2.1 p1.2 p2.2.
Instance prod_PreOrder `{Extension A} `{Extension B} `{PreOrder A ()} `{PreOrder B ()} : @PreOrder (A * B) ().
Proof.
econstructor.
- move => ?; cbv; split; by apply PreOrder_Reflexive.
- move => ? ? ? [H11 H12] [H21 H22]; split.
+ by apply (PreOrder_Transitive _ _ _ H11 H21).
+ by apply (PreOrder_Transitive _ _ _ H12 H22).
Qed.
Instance prod_PartialOrder `{Extension A} `{Extension B} `{PartialOrder A ()} `{PartialOrder B ()} : @PartialOrder (A * B) ().
Proof.
econstructor.
- by apply prod_PreOrder.
- move => [? ?] [? ?] [? ?] [? ?]. f_equal; by apply: anti_symm.
Qed.
Instance prod_Join `{Join A} `{Join B} : Join (A * B) :=
λ p1 p2, (p1.1 p2.1, p1.2 p2.2).
Instance prod_JSL `{JoinSemiLattice A} `{JoinSemiLattice B} : JoinSemiLattice (A * B).
Proof.
econstructor.
- exact: prod_PreOrder.
- move => ? ?; split; by apply join_ext_l.
- move => ? ?; split; by apply join_ext_r.
- move => ? ? ? [? ?] [? ?]. split; by apply join_least.
Qed.
Instance option_Extension `{Extension T} : Extension (option T) :=
λ o1 o2, match o1 with
| None => True
| Some t1 => match o2 with
| None => False
| Some t2 => t1 t2
end
end.
Instance option_StrictExtension `{StrictExtension T} : StrictExtension (option T) :=
λ o1 o2, match o1 with
| None => match o2 with
| None => False
| Some t2 => True
end
| Some t1 => match o2 with
| None => False
| Some t2 => t1 t2
end
end.
Instance option_StrictOrder `{StrictExtension T} `{StrictOrder T ()} : @StrictOrder (option T) ().
Proof.
econstructor.
- move => ?. cbv; repeat case_match => //. by apply: StrictOrder_Irreflexive.
- move => ? ? ?. cbv; repeat case_match => //. by apply: StrictOrder_Transitive.
Qed.
Instance option_Join `{Join T} : Join (option T) :=
λ o1 o2, match o1 with
| None => o2
| Some t1 => match o2 with
| None => o1
| Some t2 => Some $ t1 t2
end
end.
Instance option_PreOrder `{Extension T} `{PreOrder T ()} : @PreOrder (option T) ().
Proof.
econstructor.
- move => ?; cbv; repeat case_match => //.
exact: PreOrder_Reflexive.
- move => ? ? ?; cbv; repeat case_match => //.
exact: PreOrder_Transitive.
Qed.
Instance option_PartialOrder `{Extension T} `{PartialOrder T ()} : @PartialOrder (option T) ().
Proof.
econstructor.
- by eauto with typeclass_instances.
- move => [?|] [?|] //; cbv => ? ?. f_equal. exact: anti_symm.
Qed.
Instance option_Reflexive `{Extension T} `{Reflexive T ()} : @Reflexive (option T) ().
Proof. move => ?. cbv. by case_match. Qed.
Instance option_AntiSymm `{Extension T} `{AntiSymm T (=) ()} : AntiSymm (=) (() : relation (option T)).
Proof.
move => ? ?; cbv; repeat case_match => //; move => H1 H2. apply/f_equal.
by apply (@anti_symm _ _ _ H0).
Qed.
Instance option_Comm `{Join T} `{@Comm T _ (=) join} : @Comm (option T) _ (=) (join).
Proof.
move => ? ?; cbv; repeat case_match => //. f_equal. by apply comm.
Qed.
(* Instance option_Assoc `{Join T} `{@Assoc T (=) join} : @Assoc (option T) (=) (join). *)
(* Proof. *)
(* move => ? ? ?. cbv; repeat case_match => //; f_equal; simplify_option_eq. *)
(* - by apply assoc. *)
(* - reflexivity. *)
(* - reflexivity. *)
(* Qed. *)
Instance option_JSL `{JoinSemiLattice T} : JoinSemiLattice (option T).
Proof.
econstructor.
- exact: option_PreOrder.
- move => ? ?. cbv; repeat case_match => //; simplify_option_eq.
+ by apply join_ext_l.
+ by apply PreOrder_Reflexive.
- move => ? ?. cbv; repeat case_match => //; simplify_option_eq.
+ by apply join_ext_r.
+ by apply PreOrder_Reflexive.
- move => ? ? ?; cbv; repeat case_match => //; simplify_option_eq.
by apply join_least.
Qed.
(* Some auxilliary lemmas about option instances *)
Lemma option_ext_is_Some `{Extension T} (o1 o2 : option T) : o1 o2 -> is_Some o1 -> is_Some o2.
Proof.
cbv; (repeat case_match).
- by eauto.
- by eauto.
- by move => _ [] //.
Qed.
Lemma join_None `{Join T} (o1 : option T) : o1 None = o1.
Proof. by cbv; repeat case_match. Qed.
Lemma option_subseteq_None_l `{Extension T} (o1 o2 : option T) : o1 None o1 = None.
Proof. by cbv; repeat case_match. Qed.
(* Positive instances *)
Instance positive_Extension : Extension positive := ()%positive.
Instance positive_StrictExtension : StrictExtension positive := (<)%positive.
Instance positive_Join : Join positive := Pos.max.
Instance positive_PreOrder : PreOrder ()%positive.
Proof.
econstructor; unfold extension, positive_Extension;
by eauto using Pos.le_refl, Pos.le_trans.
Qed.
Instance positive_PartialOrder : PartialOrder ()%positive.
Proof.
econstructor.
- by eauto with typeclass_instances.
- exact: Pos.le_antisym.
Qed.
Instance positive_StrictOrder : StrictOrder ()%positive.
Proof.
econstructor; unfold strict_extension, positive_StrictExtension.
- move => ? ?. exact: Pos.lt_irrefl.
- exact: Pos.lt_trans.
Qed.
Instance positive_Total : Total ()%positive.
Proof.
move => x y. case: (decide (x y)%positive); try tauto.
move/Pos.lt_nle/Pos.lt_le_incl. tauto.
Qed.
Instance positive_Reflexive : Reflexive ()%positive.
Proof. exact: Pos.le_refl. Qed.
Instance positive_AntiSymm : AntiSymm (=) ()%positive.
Proof. exact: Pos.le_antisym. Qed.
(* Instance positive_Comm : Comm (=) (join)%positive. *)
(* Proof. exact: Pos.max_comm. Qed. *)
(* Instance positive_Assoc : Assoc (=) (join)%positive. *)
(* Proof. exact: Pos.max_assoc. Qed. *)
Instance poisitive_JSL : JoinSemiLattice positive.
Proof.
econstructor.
- by eauto with typeclass_instances.
- move => ? ?. exact: Pos.le_max_l.
- move => ? ?. exact: Pos.le_max_r.
- move => ? ? ?. exact: Pos.max_lub.
Qed.
(* gmap instances *)
Instance gmap_countable `{Countable K} `{Countable A} : Countable (gmap K A) :=
injective_countable (map_to_list : _ -> list (K * A))
((λ l : list (_ * _), Some (map_of_list l)) : list _ -> option _) _.
Proof.
intros. f_equal. exact: map_of_to_list.
Qed.
Instance gmap_Join `{Countable K} `{Join A} : Join (gmap K A) :=
union_with (λ a1 a2, Some $ a1 a2).
Lemma lookup_join `{Countable K} `{Join A} (f1 f2 : gmap K A) k : (f1 f2) !! k = f1 !! k f2 !! k.
Proof.
rewrite lookup_union_with.
by do 2!case: (_ !! k).
Qed.
Instance gmap_Assoc `{Countable K} `{Join A} `{@Assoc A (=) join} : @Assoc (gmap K A) (=) join.
Proof.
move => ? ? ?. apply: map_eq => k.
rewrite !lookup_union_with.
repeat case: (_ !! k) => //=.
move => ? ? ?. f_equal. by apply assoc.
Qed.
Instance gmap_Extension `{Countable K} `{Extension A} : Extension (gmap K A) :=
λ f1 f2, k, f1 !! k f2 !! k.
Instance gmap_PreOrder `{Countable K} `{Extension A} `{PreOrder A ()} : @PreOrder (gmap K A) ().
Proof.
econstructor.
- move => ? ?; reflexivity.
- move => ? ? ? ? ? ?. by etransitivity.
Qed.
Instance gmap_PartialOrder `{Countable K} `{Extension A} `{PartialOrder A ()} : @PartialOrder (gmap K A) ().
Proof.
econstructor.
- by eauto with typeclass_instances.
- move => ? ? ? ?. apply: map_eq => k. by apply: anti_symm.
Qed.
Instance gmap_JSL `{Countable K} `{JoinSemiLattice A} : JoinSemiLattice (gmap K A).
Proof.
econstructor.
- exact: gmap_PreOrder.
- move => ? ? ?. rewrite lookup_join. by apply join_ext_l.
- move => ? ? ?. rewrite lookup_join. by apply join_ext_r.
- move => ? ? ? ? ? ?. rewrite lookup_join. by apply join_least.
Qed.
(* unit instances *)
Instance unit_Extension : Extension () :=
λ u1 u2, True.
Instance unit_Join : Join () :=
λ u1 u2, tt.
Instance unit_PreOrder : @PreOrder () ().
Proof. econstructor; by auto. Qed.
Instance unit_JSL : JoinSemiLattice ().
Proof. econstructor; first exact: unit_PreOrder; by auto. Qed.
(* gset instances *)
Instance gset_Extension `{Countable K} : Extension (gset K) :=
λ s1 s2, k, k s1 k s2.
Instance gset_Join `{Countable K} : Join (gset K) :=
λ s1 s2, Mapset (join (A:=gmap K ()) (mapset_car s1) (mapset_car s2)).
Instance gset_PreOrder `{Countable K} : @PreOrder (gset K) (extension).
Proof.
econstructor.
- by move => ? ? //.
- move => ? ? ? ? ? ?. by eauto.
Qed.
Instance gset_PartialOrder `{Countable K} : @PartialOrder (gset K) (extension).
Proof.
econstructor.
- by eauto with typeclass_instances.
- move => ? ? H1 H2. apply/mapset_eq => ?. split; by [move/H1|move/H2].
Qed.
Instance gset_JSL `{Countable K} : JoinSemiLattice (gset K).
Proof.
econstructor.
- exact: gset_PreOrder.
- move => ? ? ? In. do 2!red. rewrite lookup_join. rewrite In. by case: (_ !! _) => //.
- move => ? ? ? In. do 2!red. rewrite lookup_join. rewrite In. by case: (_ !! _) => //.
- move => ? ? ? H1 H2 k.
move: (H1 k) (H2 k).
unfold elem_of, mapset_elem_of.
rewrite lookup_join. by repeat case : (_ !! _) => [[]|] //.
Qed.
Lemma gset_Assoc `{Countable K} : @Assoc (gset K) (=) join.
Proof. move => ? ? ?. by rewrite /join /gset_Join gmap_Assoc. Qed.
Lemma elem_of_join `{Countable K} (s1 s2 : gset K) (k : K) : k (join s1 s2) k s1 k s2.
Proof.
unfold elem_of, mapset_elem_of. cbn.
unfold join, gmap_Join.
rewrite lookup_union_with.
repeat case: (_ !! k) => [[]|]; tauto.
Qed.
Lemma elem_of_extension `{Countable K} (X Y : gset K) : X Y ( x : K, x X x Y).
Proof. tauto. Qed.
\ No newline at end of file
This diff is collapsed.
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