aboutsummaryrefslogtreecommitdiffstats
path: root/lib/Maps.v
diff options
context:
space:
mode:
authorxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2009-08-16 15:35:09 +0000
committerxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2009-08-16 15:35:09 +0000
commit4b119d6f9f0e846edccaf5c08788ca1615b22526 (patch)
tree66cf55decd8d950d0bdc1050448aa3ee448ca13a /lib/Maps.v
parent1fe28ba1ec3dd0657b121c4a911ee1cb046bab09 (diff)
downloadcompcert-kvx-4b119d6f9f0e846edccaf5c08788ca1615b22526.tar.gz
compcert-kvx-4b119d6f9f0e846edccaf5c08788ca1615b22526.zip
Cil2Csyntax: added goto and labels; added assignment between structs
Kildall: simplified the interface Constprop, CSE, Allocation, Linearize: updated for the new Kildall RTL, LTL: removed the well-formedness condition on the CFG, it is no longer necessary with the new Kildall and it is problematic for validated optimizations. Maps: more efficient implementation of PTree.fold. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1124 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'lib/Maps.v')
-rw-r--r--lib/Maps.v36
1 files changed, 35 insertions, 1 deletions
diff --git a/lib/Maps.v b/lib/Maps.v
index 766168ac..b607d241 100644
--- a/lib/Maps.v
+++ b/lib/Maps.v
@@ -887,15 +887,49 @@ Module PTree <: TREE.
intros. change (list_norepet (xkeys m 1)). apply xelements_keys_norepet.
Qed.
+(*
Definition fold (A B : Type) (f: B -> positive -> A -> B) (tr: t A) (v: B) :=
List.fold_left (fun a p => f a (fst p) (snd p)) (elements tr) v.
+*)
+
+ Fixpoint xfold (A B: Type) (f: B -> positive -> A -> B)
+ (i: positive) (m: t A) (v: B) {struct m} : B :=
+ match m with
+ | Leaf => v
+ | Node l None r =>
+ let v1 := xfold f (append i (xO xH)) l v in
+ xfold f (append i (xI xH)) r v1
+ | Node l (Some x) r =>
+ let v1 := xfold f (append i (xO xH)) l v in
+ let v2 := f v1 i x in
+ xfold f (append i (xI xH)) r v2
+ end.
+
+ Definition fold (A B : Type) (f: B -> positive -> A -> B) (m: t A) (v: B) :=
+ xfold f xH m v.
+
+ Lemma xfold_xelements:
+ forall (A B: Type) (f: B -> positive -> A -> B) m i v,
+ xfold f i m v =
+ List.fold_left (fun a p => f a (fst p) (snd p))
+ (xelements m i)
+ v.
+ Proof.
+ induction m; intros.
+ simpl. auto.
+ simpl. destruct o.
+ rewrite fold_left_app. simpl.
+ rewrite IHm1. apply IHm2.
+ rewrite fold_left_app. simpl.
+ rewrite IHm1. apply IHm2.
+ Qed.
Theorem fold_spec:
forall (A B: Type) (f: B -> positive -> A -> B) (v: B) (m: t A),
fold f m v =
List.fold_left (fun a p => f a (fst p) (snd p)) (elements m) v.
Proof.
- intros; unfold fold; auto.
+ intros. unfold fold, elements. apply xfold_xelements.
Qed.
End PTree.