Skip to content

Commit

Permalink
Adapt w.r.t. coq/coq#16004
Browse files Browse the repository at this point in the history
  • Loading branch information
Alizter committed May 19, 2022
1 parent 0fd5faa commit 8e964f5
Show file tree
Hide file tree
Showing 24 changed files with 58 additions and 58 deletions.
2 changes: 1 addition & 1 deletion src/riscv/Examples/Example64Literal.v
Original file line number Diff line number Diff line change
Expand Up @@ -19,7 +19,7 @@ Require Import riscv.Utility.DefaultMemImpl64.
Require Import coqutil.Map.Z_keyed_SortedListMap.
Require coqutil.Map.SortedList.

Existing Instance DefaultRiscvState.
#[global] Existing Instance DefaultRiscvState.


Definition foo7prog: list MachineInt := [
Expand Down
2 changes: 1 addition & 1 deletion src/riscv/Examples/MulTrapHandler.v
Original file line number Diff line number Diff line change
Expand Up @@ -73,7 +73,7 @@ Definition initial_state: State := HNil
(CSRField.MScratch, handler_stack_start) :: nil)]
[exectrace := nil].

Instance IsRiscvMachine: RiscvProgram (StateAbortFail State) word :=
#[global] Instance IsRiscvMachine: RiscvProgram (StateAbortFail State) word :=
AddExecTrace FieldNames (MinimalCSRsDet.IsRiscvMachine FieldNames).

(* success flag * final state *)
Expand Down
2 changes: 1 addition & 1 deletion src/riscv/Examples/SMTVerif.v
Original file line number Diff line number Diff line change
Expand Up @@ -59,7 +59,7 @@ Definition loadN(n: nat)(kind: SourceType)(a: word): OState MachineState (HList.
Definition storeN(n: nat)(kind: SourceType)(a: word)(v: HList.tuple byte n): OState MachineState unit :=
fail_hard.

Instance IsRiscvProgram: RiscvProgram (OState MachineState) word := {
#[global] Instance IsRiscvProgram: RiscvProgram (OState MachineState) word := {
getRegister reg :=
if Z.eq_dec reg Register0 then
Return (ZToReg 0)
Expand Down
4 changes: 2 additions & 2 deletions src/riscv/Examples/WMMFree.v
Original file line number Diff line number Diff line change
Expand Up @@ -278,7 +278,7 @@ Fixpoint bind{A B: Type}(m: M A){struct m}: (A -> M B) -> M B :=
| HaltThread => fun f => HaltThread (* Note: Drops the continuation in f *)
end.

Instance M_Monad: Monad M. refine ({|
#[global] Instance M_Monad: Monad M. refine ({|
Bind := @bind;
Return := @Ret
|}).
Expand Down Expand Up @@ -471,7 +471,7 @@ Definition updateDeps(inst: Instruction): Array Register (set Event) -> Array Re
| _ => id (* only I is supported at the moment *)
end.

Instance IsRiscvMachine: RiscvProgram M word := {
#[global] Instance IsRiscvMachine: RiscvProgram M word := {
getRegister reg := s <- get; Return (getReg s.(Regs) reg);
setRegister reg v := s <- get; put (withRegs (setReg s.(Regs) reg v) s);
getPC := s <- get; Return s.(Pc);
Expand Down
2 changes: 1 addition & 1 deletion src/riscv/Platform/AtomicMinimal.v
Original file line number Diff line number Diff line change
Expand Up @@ -73,4 +73,4 @@ Section Riscv.

End Riscv.

Existing Instance IsAtomicRiscvMachine.
#[global] Existing Instance IsAtomicRiscvMachine.
2 changes: 1 addition & 1 deletion src/riscv/Platform/FE310ExtSpec.v
Original file line number Diff line number Diff line change
Expand Up @@ -24,4 +24,4 @@ Section MMIO.

End MMIO.

Hint Resolve FE310_mmio: typeclass_instances.
#[global] Hint Resolve FE310_mmio: typeclass_instances.
2 changes: 1 addition & 1 deletion src/riscv/Platform/MetricLogging.v
Original file line number Diff line number Diff line change
Expand Up @@ -59,7 +59,7 @@ Delimit Scope MetricL_scope with metricsL.
Infix "<=" := metricsLeq : MetricL_scope.
Infix "-" := metricsSub : MetricL_scope.

Hint Unfold
#[global] Hint Unfold
withInstructions
withLoads
withStores
Expand Down
2 changes: 1 addition & 1 deletion src/riscv/Platform/MetricMinimal.v
Original file line number Diff line number Diff line change
Expand Up @@ -151,4 +151,4 @@ Section Riscv.

End Riscv.

Existing Instance IsMetricRiscvMachine.
#[global] Existing Instance IsMetricRiscvMachine.
4 changes: 2 additions & 2 deletions src/riscv/Platform/Minimal.v
Original file line number Diff line number Diff line change
Expand Up @@ -260,5 +260,5 @@ Section Riscv.
End Riscv.

(* needed because defined inside a Section *)
Existing Instance IsRiscvMachine.
Existing Instance MinimalSatisfiesPrimitives.
#[global] Existing Instance IsRiscvMachine.
#[global] Existing Instance MinimalSatisfiesPrimitives.
2 changes: 1 addition & 1 deletion src/riscv/Platform/MinimalCSRsDet.v
Original file line number Diff line number Diff line change
Expand Up @@ -120,4 +120,4 @@ Section Riscv.
End Riscv.

(* needed because defined inside a Section *)
Existing Instance IsRiscvMachine.
#[global] Existing Instance IsRiscvMachine.
2 changes: 1 addition & 1 deletion src/riscv/Platform/MinimalLogging.v
Original file line number Diff line number Diff line change
Expand Up @@ -77,4 +77,4 @@ Section Riscv.

End Riscv.

Existing Instance IsRiscvMachineL. (* needed because it was defined inside a Section *)
#[global] Existing Instance IsRiscvMachineL. (* needed because it was defined inside a Section *)
2 changes: 1 addition & 1 deletion src/riscv/Proofs/EncodeBound.v
Original file line number Diff line number Diff line change
Expand Up @@ -68,7 +68,7 @@ Lemma encode_U_bitSlice_idemp: forall opcode r z,
encode_U opcode r z.
Proof. intros. unfold encode_U. prove_Zeq_bitwise. Qed.

Hint Rewrite
#[global] Hint Rewrite
encode_Fence_bitSlice_idemp
encode_Invalid_bitSlice_idemp
encode_I_bitSlice_idemp
Expand Down
8 changes: 4 additions & 4 deletions src/riscv/Spec/CSRFile.v
Original file line number Diff line number Diff line change
Expand Up @@ -87,7 +87,7 @@ Proof. intros. destruct f1; destruct f2; simpl in *; congruence. Qed.

Definition CSRField_ltb(f1 f2: CSRField): bool := Z.ltb (CSRField_to_Z f1) (CSRField_to_Z f2).

Instance CSRField_ltb_strictorder: SortedList.parameters.strict_order CSRField_ltb.
#[global] Instance CSRField_ltb_strictorder: SortedList.parameters.strict_order CSRField_ltb.
Proof.
constructor; unfold CSRField_ltb; intros.
- apply Z.ltb_irrefl.
Expand All @@ -97,14 +97,14 @@ Proof.
blia.
Qed.

Instance CSRFile_map_params: SortedList.parameters := {|
#[global] Instance CSRFile_map_params: SortedList.parameters := {|
parameters.key := CSRField;
parameters.value := MachineInt;
parameters.ltb := CSRField_ltb;
|}.

Instance CSRFile: map.map CSRField MachineInt :=
#[global] Instance CSRFile: map.map CSRField MachineInt :=
SortedList.map CSRFile_map_params CSRField_ltb_strictorder.

Instance CSRFile_ok: map.ok CSRFile.
#[global] Instance CSRFile_ok: map.ok CSRFile.
Proof. apply (@SortedList.map_ok CSRFile_map_params CSRField_ltb_strictorder). Qed.
2 changes: 1 addition & 1 deletion src/riscv/Spec/Machine.v
Original file line number Diff line number Diff line change
Expand Up @@ -147,4 +147,4 @@ Arguments RiscvProgram (M) (t) {_} {_}.
Arguments RiscvMachine: clear implicits.
Arguments RiscvMachine (M) (t) {_} {_} {_}.

Existing Instance DefaultRiscvState.
#[global] Existing Instance DefaultRiscvState.
6 changes: 3 additions & 3 deletions src/riscv/Utility/DefaultMemImpl32.v
Original file line number Diff line number Diff line change
Expand Up @@ -6,11 +6,11 @@ Require Import riscv.Utility.Words32Naive.
Require coqutil.Map.SortedList coqutil.Map.SortedListWord.


Instance params: SortedList.parameters := {|
#[global] Instance params: SortedList.parameters := {|
SortedList.parameters.key := word32;
SortedList.parameters.value := Byte.byte;
SortedList.parameters.ltb := word.ltu;
|}.

Instance Mem: map.map word32 Byte.byte := SortedList.map params _.
Instance MemOk: map.ok Mem := SortedList.map_ok.
#[global] Instance Mem: map.map word32 Byte.byte := SortedList.map params _.
#[global] Instance MemOk: map.ok Mem := SortedList.map_ok.
6 changes: 3 additions & 3 deletions src/riscv/Utility/DefaultMemImpl64.v
Original file line number Diff line number Diff line change
Expand Up @@ -6,11 +6,11 @@ Require Import riscv.Utility.Words64Naive.
Require coqutil.Map.SortedList coqutil.Map.SortedListWord.


Instance params: SortedList.parameters := {|
#[global] Instance params: SortedList.parameters := {|
SortedList.parameters.key := word64;
SortedList.parameters.value := Byte.byte;
SortedList.parameters.ltb := word.ltu;
|}.

Instance Mem: map.map word64 Byte.byte := SortedList.map params _.
Instance MemOk: map.ok Mem := SortedList.map_ok.
#[global] Instance Mem: map.map word64 Byte.byte := SortedList.map params _.
#[global] Instance MemOk: map.ok Mem := SortedList.map_ok.
2 changes: 1 addition & 1 deletion src/riscv/Utility/Encode.v
Original file line number Diff line number Diff line change
Expand Up @@ -418,7 +418,7 @@ Definition verify_iset(inst: Instruction)(iset: InstructionSet): Prop :=
Definition verify(inst: Instruction)(iset: InstructionSet): Prop :=
respects_bounds (bitwidth iset) inst /\ verify_iset inst iset.

Hint Unfold
#[global] Hint Unfold
map_Invalid
map_R
map_R_atomic
Expand Down
10 changes: 5 additions & 5 deletions src/riscv/Utility/JMonad.v
Original file line number Diff line number Diff line change
Expand Up @@ -21,7 +21,7 @@ Open Scope monad_scope.

Definition Id: Type -> Type := id.

Instance Id_monad: Monad Id := {|
#[global] Instance Id_monad: Monad Id := {|
ret := @id;
mmap A B := @id (A -> B);
join := @id;
Expand All @@ -48,7 +48,7 @@ Definition flattenOption{A}(ooa: option (option A)): option A :=
| _ => None
end.

Instance OptionT_Monad(M: Type -> Type){MM: Monad M}: Monad (optionT M) := {|
#[global] Instance OptionT_Monad(M: Type -> Type){MM: Monad M}: Monad (optionT M) := {|
ret A (a: A) := mkOptionT (ret (retOption a));
mmap A B (f: A -> B)(oma: optionT M A) :=
mkOptionT (mmap (mapOption f)
Expand All @@ -75,7 +75,7 @@ Record listT(M: Type -> Type)(A: Type): Type := mkListT {
Arguments mkListT {M} {A} (_).
Arguments runListT {M} {A} (_).

Instance listT_Monad(M: Type -> Type){MM: Monad M}: Monad (listT M) := {|
#[global] Instance listT_Monad(M: Type -> Type){MM: Monad M}: Monad (listT M) := {|
ret A (a: A) := mkListT (ret (retList a));
mmap A B (f: A -> B)(nma: listT M A) :=
mkListT (mmap (mapList f)
Expand All @@ -98,7 +98,7 @@ Record StateT(S: Type)(M: Type -> Type)(A: Type): Type := mkStateT {
Arguments mkStateT {S} {M} {A} (_).
Arguments runStateT {S} {M} {A} (_).

Instance StateT_Monad(S: Type)(M: Type -> Type){MM: Monad M}: Monad (StateT S M) := {|
#[global] Instance StateT_Monad(S: Type)(M: Type -> Type){MM: Monad M}: Monad (StateT S M) := {|
ret A (a: A) := mkStateT (fun s => ret (a, s));
mmap A B (f: A -> B)(sma: StateT S M A) :=
mkStateT (fun s => mmap (fun '(a, s0) => (f a, s0)) (runStateT sma s));
Expand Down Expand Up @@ -171,7 +171,7 @@ Instance NDStateT_Monad(S: Type)(M: Type -> Type){MM: Monad M}: Monad (NDStateT
*)

(* note: no "pick" is needed as argument here *)
Instance NDS_Monad(S: Type): Monad (NDS S) := {|
#[global] Instance NDS_Monad(S: Type): Monad (NDS S) := {|
ret A (a: A) :=
fun s => cons (a, s) nil;
mmap A B (f: A -> B)(m: NDS S A) :=
Expand Down
2 changes: 1 addition & 1 deletion src/riscv/Utility/MkMachineWidth.v
Original file line number Diff line number Diff line change
Expand Up @@ -6,7 +6,7 @@ Require Import coqutil.Datatypes.HList.
Require Import riscv.Utility.Utility.
Local Open Scope Z_scope.

Instance MachineWidth_XLEN{width}{_: Bitwidth width}{word: word width}: MachineWidth word := {|
#[global] Instance MachineWidth_XLEN{width}{_: Bitwidth width}{word: word width}: MachineWidth word := {|
add := word.add;
sub := word.sub;
mul := word.mul;
Expand Down
12 changes: 6 additions & 6 deletions src/riscv/Utility/MonadT.v
Original file line number Diff line number Diff line change
Expand Up @@ -21,7 +21,7 @@ Notation "m1 ;; m2" := (Bind m1 (fun _ => m2))

Open Scope monad_scope.

Instance option_Monad: Monad option := {|
#[global] Instance option_Monad: Monad option := {|
Bind := fun {A B: Type} (o: option A) (f: A -> option B) => match o with
| Some x => f x
| None => None
Expand Down Expand Up @@ -75,7 +75,7 @@ listT option nat
(*Eval cbv in (optionT list nat).*)


Instance OptionT_Monad(M: Type -> Type){MM: Monad M}: Monad (optionT M) := {|
#[global] Instance OptionT_Monad(M: Type -> Type){MM: Monad M}: Monad (optionT M) := {|
Bind{A}{B}(m: M (option A))(f: A -> M (option B)) :=
Bind m (fun (o: option A) =>
match o with
Expand All @@ -98,7 +98,7 @@ End OptionMonad.

Definition State(S A: Type) := S -> (A * S).

Instance State_Monad(S: Type): Monad (State S) := {|
#[global] Instance State_Monad(S: Type): Monad (State S) := {|
Bind := fun {A B: Type} (m: State S A) (f: A -> State S B) =>
fun (s: S) => let (a, s') := m s in f a s' ;
Return := fun {A: Type} (a: A) =>
Expand All @@ -107,7 +107,7 @@ Instance State_Monad(S: Type): Monad (State S) := {|

Definition StateT(S: Type)(M: Type -> Type)(A: Type) := S -> M (A * S)%type.

Instance StateT_Monad(M: Type -> Type){MM: Monad M}(S: Type): Monad (StateT S M) := {|
#[global] Instance StateT_Monad(M: Type -> Type){MM: Monad M}(S: Type): Monad (StateT S M) := {|
Bind{A B: Type}(m: StateT S M A)(f: A -> StateT S M B) :=
fun (s: S) => Bind (m s) (fun '(a, s) => f a s);
Return{A: Type}(a: A) :=
Expand Down Expand Up @@ -174,7 +174,7 @@ Module NonDetMonad.

End NonDetMonad.

Instance NonDet_Monad: Monad NonDet := {|
#[global] Instance NonDet_Monad: Monad NonDet := {|
Bind := @NonDetMonad.flatMapSet;
Return := @NonDetMonad.singletonSet;
|}.
Expand All @@ -184,7 +184,7 @@ Definition NonDetT(M: Type -> Type)(A: Type) := NonDet (M A).
Goal forall (M: Type -> Type) (A: Type), NonDet (M A) = (M A -> Prop).
intros. reflexivity. Qed.

Instance NonDetT_Monad(M: Type -> Type){MM: Monad M}: Monad (NonDetT M). refine ({|
#[global] Instance NonDetT_Monad(M: Type -> Type){MM: Monad M}: Monad (NonDetT M). refine ({|
Bind{A B}(m: NonDet (M A))(f: A -> NonDet (M B)) := _;
Return := _
|}).
Expand Down
10 changes: 5 additions & 5 deletions src/riscv/Utility/MonadTests.v
Original file line number Diff line number Diff line change
Expand Up @@ -17,15 +17,15 @@ Open Scope monad_scope.

Definition Id: Type -> Type := id.

Instance Id_monad: Monad Id := {|
#[global] Instance Id_monad: Monad Id := {|
Bind{A B}(m: A)(f: A -> B) := f m;
Return{A}(a: A) := a;
|}.


Definition NonDet(A: Type): Type := A -> Prop.

Instance NonDet_Monad: Monad NonDet := {|
#[global] Instance NonDet_Monad: Monad NonDet := {|
Bind{A B}(m: NonDet A)(f: A -> NonDet B) :=
fun (b: B) => exists a, m a /\ f a b;
Return{A} := eq;
Expand All @@ -38,7 +38,7 @@ Record optionT(M: Type -> Type)(A: Type): Type := mkOptionT {
Arguments mkOptionT {M} {A} (_).
Arguments runOptionT {M} {A} (_).

Instance OptionT_Monad(M: Type -> Type){MM: Monad M}: Monad (optionT M) := {|
#[global] Instance OptionT_Monad(M: Type -> Type){MM: Monad M}: Monad (optionT M) := {|
Bind{A}{B}(m: optionT M A)(f: A -> optionT M B) :=
mkOptionT (Bind (runOptionT m) (fun (o: option A) =>
match o with
Expand All @@ -63,7 +63,7 @@ Record StateT(S: Type)(M: Type -> Type)(A: Type): Type := mkStateT {
Arguments mkStateT {S} {M} {A} (_).
Arguments runStateT {S} {M} {A} (_).

Instance StateT_Monad(M: Type -> Type){MM: Monad M}(S: Type): Monad (StateT S M) := {|
#[global] Instance StateT_Monad(M: Type -> Type){MM: Monad M}(S: Type): Monad (StateT S M) := {|
Bind{A B: Type}(m: StateT S M A)(f: A -> StateT S M B) :=
mkStateT (fun (s: S) => Bind ((runStateT m) s) (fun '(a, s) => runStateT (f a) s));
Return{A: Type}(a: A) :=
Expand All @@ -86,7 +86,7 @@ Record listT(M: Type -> Type)(A: Type): Type := mkListT {
Arguments mkListT {M} {A} (_).
Arguments runListT {M} {A} (_).

Instance listT_Monad(M: Type -> Type){MM: Monad M}: Monad (listT M) := {|
#[global] Instance listT_Monad(M: Type -> Type){MM: Monad M}: Monad (listT M) := {|
Bind{A B: Type}(m: listT M A)(f: A -> listT M B) :=
mkListT (la <- runListT m;
List.fold_left
Expand Down
Loading

0 comments on commit 8e964f5

Please sign in to comment.