Lvc.IL.Var
Require Import List.
Require Import Util EqDec.
Require Import OrderedTypeEx.
Set Implicit Arguments.
Require Import Util EqDec.
Require Import OrderedTypeEx.
Set Implicit Arguments.
Definitions related to variables
A class for counted types
A counted type is a type for which an injection to the naturals exists
Class Counted (A : Type) := {
counted : A → nat;
counted_inj : ∀ (x y : A), counted x = counted y → x = y;
incc : ∀ (x : A), nat → A
}.
counted : A → nat;
counted_inj : ∀ (x y : A), counted x = counted y → x = y;
incc : ∀ (x : A), nat → A
}.
Definition var : Type := nat.
Definition default_var : var := 0%nat.
Global Instance inst_defaulted_var : Defaulted var := Build_Defaulted default_var.
Global Instance inst_eq_dec_var : EqDec var eq := nat_eq_eqdec.
Global Program Instance inst_counted_var : Counted var :=
Build_Counted (fun x ⇒ x) _ (fun x ⇒ fun y ⇒ x + y).
Definition default_var : var := 0%nat.
Global Instance inst_defaulted_var : Defaulted var := Build_Defaulted default_var.
Global Instance inst_eq_dec_var : EqDec var eq := nat_eq_eqdec.
Global Program Instance inst_counted_var : Counted var :=
Build_Counted (fun x ⇒ x) _ (fun x ⇒ fun y ⇒ x + y).
Inductive loc : Type :=
| LocI : nat → loc.
Definition default_loc : loc := LocI 0.
Global Instance inst_eq_dec_loc : EqDec loc eq.
Defined.
Definition locN l :=
match l with
| LocI n ⇒ n
end.
Lemma locN_inj (x y : loc)
: locN x = locN y → x = y.
Definition locInc (l:loc) (d:nat) := match l with LocI n ⇒ LocI (d + n) end.
Global Instance inst_counted_loc : Counted loc :=
Build_Counted locN locN_inj locInc.
Inductive lab : Type :=
| LabI : nat → lab.
Definition default_lab := LabI 0.
Lemma eq_dec_lab (x y : lab) :
{x=y} + {x≠y}.
| LabI : nat → lab.
Definition default_lab := LabI 0.
Lemma eq_dec_lab (x y : lab) :
{x=y} + {x≠y}.
Definition labN (x : lab) : nat :=
match x with
| LabI x ⇒ x
end.
Lemma labN_inj (x y : lab)
: labN x = labN y → x = y.
Definition labInc (l:lab) (d:nat) := match l with LabI n ⇒ LabI (d + n) end.
Global Instance inst_counted_lab : Counted lab :=
Build_Counted labN labN_inj labInc.
Lemma counted_labI (n : nat) :
counted (LabI n) = n.
Definition lt (l l´:lab) :=
match l, l´ with
LabI n, LabI m ⇒ lt n m
end.
match x with
| LabI x ⇒ x
end.
Lemma labN_inj (x y : lab)
: labN x = labN y → x = y.
Definition labInc (l:lab) (d:nat) := match l with LabI n ⇒ LabI (d + n) end.
Global Instance inst_counted_lab : Counted lab :=
Build_Counted labN labN_inj labInc.
Lemma counted_labI (n : nat) :
counted (LabI n) = n.
Definition lt (l l´:lab) :=
match l, l´ with
LabI n, LabI m ⇒ lt n m
end.
Instance lab_StrictOrder : StrictOrder lt (@eq lab).
Qed.
Fixpoint lab_compare (l l´ : lab) :=
match l, l´ with
| LabI n, LabI m ⇒ nat_compare n m
end.
Program Instance lab_OrderedType : UsualOrderedType lab := {
SOT_lt := lt;
SOT_cmp := lab_compare;
SOT_StrictOrder := lab_StrictOrder
}.
Instance proper_var (ϱ:var → var)
: Proper (_eq ==> _eq) ϱ.
Hint Extern 20 (@Proper
(@respectful _ _
(@_eq _ (@SOT_as_OT _ (@eq nat) nat_OrderedType))
(@_eq _ (@SOT_as_OT _ (@eq nat) nat_OrderedType))) ?ϱ) ⇒ eapply proper_var.
Qed.
Fixpoint lab_compare (l l´ : lab) :=
match l, l´ with
| LabI n, LabI m ⇒ nat_compare n m
end.
Program Instance lab_OrderedType : UsualOrderedType lab := {
SOT_lt := lt;
SOT_cmp := lab_compare;
SOT_StrictOrder := lab_StrictOrder
}.
Instance proper_var (ϱ:var → var)
: Proper (_eq ==> _eq) ϱ.
Hint Extern 20 (@Proper
(@respectful _ _
(@_eq _ (@SOT_as_OT _ (@eq nat) nat_OrderedType))
(@_eq _ (@SOT_as_OT _ (@eq nat) nat_OrderedType))) ?ϱ) ⇒ eapply proper_var.