(**************************************************************************) (* *) (* SMTCoq *) (* Copyright (C) 2011 - 2022 *) (* *) (* See file "AUTHORS" for the list of authors *) (* *) (* This file is distributed under the terms of the CeCILL-C licence *) (* *) (**************************************************************************) (* Software implementation of arrays, based on finite maps using AVL trees *) Declare Scope array_scope. Require Export Int63 Psatz. Require FMapAVL. Require Import ZArith. Local Open Scope int63_scope. Import OrderedType. Module IntOrderedType <: OrderedType. Definition t := int. Definition eq x y := (x =? y) = true. Definition lt x y := (x eq y x. Proof. unfold eq. rewrite !eqb_spec. intros ->. reflexivity. Qed. Lemma eq_trans x y z : eq x y -> eq y z -> eq x z. Proof. unfold eq. rewrite !eqb_spec. intros -> ->. reflexivity. Qed. Lemma lt_trans x y z : lt x y -> lt y z -> lt x z. Proof. unfold lt; rewrite !ltb_spec; apply Z.lt_trans. Qed. Lemma lt_not_eq x y : lt x y -> ~ eq x y. Proof. unfold lt, eq. rewrite ltb_spec, eqb_spec. intros H1 H2. rewrite H2 in H1. lia. Qed. Definition compare x y : Compare lt eq x y. Proof. case_eq (x x | None => d end else d. Definition default {A:Type} (t:array A) : A := let (td,_) := t in let (_,d) := td in d. Definition set {A:Type} (t:array A) (i:int) (a:A) : array A := let (td,l) := t in if l <=? i then t else let (t,d) := td in (Map.add i a t, d, l). Definition length {A:Type} (t:array A) : int := let (_,l) := t in l. Definition copy {A:Type} (t:array A) : array A := t. Module Export PArrayNotations. Delimit Scope array_scope with array. Notation "t '.[' i ']'" := (get t i) (at level 50) : array_scope. Notation "t '.[' i '<-' a ']'" := (set t i a) (at level 50) : array_scope. End PArrayNotations. Local Open Scope array_scope. Definition max_length := max_int. (** Axioms *) Require FSets.FMapFacts. Module P := FSets.FMapFacts.WProperties_fun IntOrderedType Map. Lemma get_outofbound : forall A (t:array A) i, (i t.[i] = default t. intros A t i; unfold get. destruct t as ((t, d), l). unfold length; intro Hi; rewrite Hi; clear Hi. reflexivity. Qed. Lemma get_set_same : forall A t i (a:A), (i t.[i<-a].[i] = a. intros A t i a. destruct t as ((t, d), l). unfold set, get, length. intro Hi; generalize Hi. rewrite ltb_spec. rewrite Z.lt_nge. rewrite <- leb_spec. rewrite Bool.not_true_iff_false. intro Hi'; rewrite Hi'; clear Hi'. rewrite Hi; clear Hi. rewrite P.F.add_eq_o. reflexivity. rewrite eqb_spec. reflexivity. Qed. Lemma get_set_other : forall A t i j (a:A), i <> j -> t.[i<-a].[j] = t.[j]. intros A t i j a Hij. destruct t as ((t, d), l). unfold set, get, length. case (l <=? i). reflexivity. rewrite P.F.add_neq_o. reflexivity. intro H; apply Hij; clear Hij. rewrite eqb_spec in H. assumption. Qed. Lemma default_set : forall A t i (a:A), default (t.[i<-a]) = default t. intros A t i a. destruct t as ((t, d), l). unfold default, set. case (l <=? i); reflexivity. Qed. Lemma get_make : forall A (a:A) size i, (make size a).[i] = a. intros A a size i. unfold make, get. rewrite P.F.empty_o. case (i (forall i, i t1.[i] = t2.[i]) -> default t1 = default t2 -> t1 = t2. *) (* Lemmas *) Lemma default_copy A (t:array A) : default (copy t) = default t. unfold copy; reflexivity. Qed. Lemma default_make A (a : A) size : default (make size a) = a. unfold default, make; reflexivity. Qed. Lemma get_set_same_default A (t : array A) (i : int) : t.[i <- default t].[i] = default t. unfold default, get, set. destruct t as ((t, d), l). case_eq (i default t -> (x