aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-09-03 11:53:59 +0200
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-09-03 11:53:59 +0200
commit38e8be31f6445b42526ad577bc388f821649b062 (patch)
tree533656e6c1da78581f433239d705d44f446dc93c
parentb70d80e7259873ac941830e02b022ca8e92541a6 (diff)
parent7ae2140a2763f1e646630ceca27c2088aa31cf00 (diff)
downloadcompcert-kvx-38e8be31f6445b42526ad577bc388f821649b062.tar.gz
compcert-kvx-38e8be31f6445b42526ad577bc388f821649b062.zip
Merge remote-tracking branch 'origin/mppa-work' into mppa-non-trapping-load
-rw-r--r--LICENSE27
-rw-r--r--Makefile25
-rw-r--r--Makefile.extr5
-rw-r--r--MenhirLib/Alphabet.v247
-rw-r--r--MenhirLib/Automaton.v (renamed from cparser/MenhirLib/Automaton.v)37
-rw-r--r--MenhirLib/Grammar.v162
-rw-r--r--MenhirLib/Interpreter.v453
-rw-r--r--MenhirLib/Interpreter_complete.v825
-rw-r--r--MenhirLib/Interpreter_correct.v175
-rw-r--r--MenhirLib/Main.v79
-rw-r--r--MenhirLib/Validator_classes.v75
-rw-r--r--MenhirLib/Validator_complete.v394
-rw-r--r--MenhirLib/Validator_safe.v234
-rw-r--r--arm/Archi.v61
-rw-r--r--arm/Asmexpand.ml5
-rw-r--r--arm/Builtins1.v (renamed from cparser/Builtins.mli)26
-rw-r--r--arm/CBuiltins.ml4
-rw-r--r--arm/ConstpropOp.vp12
-rw-r--r--arm/ConstpropOpproof.v26
-rw-r--r--arm/SelectOp.vp13
-rw-r--r--arm/SelectOpproof.v27
-rw-r--r--arm/TargetPrinter.ml6
-rw-r--r--backend/CSE.v15
-rw-r--r--backend/CSEproof.v6
-rw-r--r--backend/Constprop.v44
-rw-r--r--backend/Constpropproof.v44
-rw-r--r--backend/Conventions.v2
-rw-r--r--backend/Linearizeaux.ml2
-rw-r--r--backend/OpHelpersproof.v2
-rw-r--r--backend/PrintAsm.ml3
-rw-r--r--backend/PrintAsmaux.ml23
-rw-r--r--backend/PrintCminor.ml2
-rw-r--r--backend/PrintLTL.ml2
-rw-r--r--backend/PrintRTL.ml2
-rw-r--r--backend/PrintXTL.ml2
-rw-r--r--backend/Selection.v56
-rw-r--r--backend/Selectionaux.ml11
-rw-r--r--backend/Selectionproof.v157
-rw-r--r--backend/SplitLongproof.v104
-rw-r--r--backend/ValueAnalysis.v52
-rw-r--r--backend/ValueDomain.v44
-rw-r--r--cfrontend/C2C.ml74
-rw-r--r--cfrontend/Cexec.v54
-rw-r--r--cfrontend/Cop.v25
-rw-r--r--cfrontend/Csem.v69
-rw-r--r--cfrontend/Cstrategy.v10
-rw-r--r--cfrontend/Csyntax.v12
-rw-r--r--cfrontend/Ctyping.v77
-rw-r--r--cfrontend/SimplExprspec.v2
-rw-r--r--common/Builtins.v58
-rw-r--r--common/Builtins0.v531
-rw-r--r--common/Events.v88
-rw-r--r--common/Separation.v2
-rw-r--r--common/Values.v10
-rwxr-xr-xconfigure24
-rw-r--r--cparser/Builtins.ml54
-rw-r--r--cparser/C.mli7
-rw-r--r--cparser/Cabs.v54
-rw-r--r--cparser/Cabshelper.ml13
-rw-r--r--cparser/Checks.ml362
-rw-r--r--cparser/Checks.mli2
-rw-r--r--cparser/Cutil.ml7
-rw-r--r--cparser/Cutil.mli2
-rw-r--r--cparser/Diagnostics.ml8
-rw-r--r--cparser/Diagnostics.mli4
-rw-r--r--cparser/Elab.ml43
-rw-r--r--cparser/Elab.mli4
-rw-r--r--cparser/Env.ml40
-rw-r--r--cparser/Env.mli7
-rw-r--r--cparser/GCC.ml4
-rw-r--r--cparser/Lexer.mll247
-rw-r--r--cparser/MenhirLib/Alphabet.v320
-rw-r--r--cparser/MenhirLib/Grammar.v166
-rw-r--r--cparser/MenhirLib/Interpreter.v228
-rw-r--r--cparser/MenhirLib/Interpreter_complete.v686
-rw-r--r--cparser/MenhirLib/Interpreter_correct.v228
-rw-r--r--cparser/MenhirLib/Interpreter_safe.v275
-rw-r--r--cparser/MenhirLib/Main.v92
-rw-r--r--cparser/MenhirLib/Tuples.v49
-rw-r--r--cparser/MenhirLib/Validator_complete.v542
-rw-r--r--cparser/MenhirLib/Validator_safe.v414
-rw-r--r--cparser/PackedStructs.ml2
-rw-r--r--cparser/Parse.ml13
-rw-r--r--cparser/Parser.vy595
-rw-r--r--cparser/Rename.ml2
-rw-r--r--cparser/Transform.ml2
-rw-r--r--cparser/Unblock.ml2
-rw-r--r--cparser/pre_parser.mly10
-rw-r--r--debug/DebugInformation.ml2
-rw-r--r--debug/DwarfPrinter.ml4
-rw-r--r--doc/ccomp.120
-rw-r--r--driver/Clflags.ml2
-rw-r--r--driver/Driver.ml6
-rw-r--r--driver/Frontend.ml2
-rw-r--r--driver/Interp.ml4
-rw-r--r--exportclight/Clightgen.ml33
-rw-r--r--extraction/extraction.v15
-rw-r--r--flocq/IEEE754/Binary.v121
-rw-r--r--lib/Camlcoq.ml51
-rw-r--r--lib/Floats.v316
-rw-r--r--lib/Maps.v2
-rw-r--r--mppa_k1c/Archi.v37
-rw-r--r--mppa_k1c/Asm.v20
-rw-r--r--mppa_k1c/Asmblockdeps.v9
-rw-r--r--mppa_k1c/Asmblockgen.v40
-rw-r--r--mppa_k1c/Asmexpand.ml4
-rw-r--r--mppa_k1c/Asmvliw.v19
-rw-r--r--mppa_k1c/Builtins1.v66
-rw-r--r--mppa_k1c/CBuiltins.ml23
-rw-r--r--mppa_k1c/ExtFloats.v39
-rw-r--r--mppa_k1c/ExtValues.v49
-rw-r--r--mppa_k1c/Machregs.v2
-rw-r--r--mppa_k1c/NeedOp.v60
-rw-r--r--mppa_k1c/Op.v73
-rw-r--r--mppa_k1c/PostpassSchedulingOracle.ml22
-rw-r--r--mppa_k1c/SelectOp.vp41
-rw-r--r--mppa_k1c/SelectOpproof.v103
-rw-r--r--mppa_k1c/TargetPrinter.ml22
-rw-r--r--mppa_k1c/ValueAOp.v134
-rw-r--r--mppa_k1c/lib/Machblock.v4
-rw-r--r--mppa_k1c/lib/Machblockgen.v13
-rwxr-xr-xpg20
-rw-r--r--powerpc/Archi.v35
-rw-r--r--powerpc/Asmexpand.ml20
-rw-r--r--powerpc/Builtins1.v33
-rw-r--r--powerpc/CBuiltins.ml4
-rw-r--r--powerpc/ConstpropOp.vp13
-rw-r--r--powerpc/ConstpropOpproof.v28
-rw-r--r--powerpc/SelectOp.vp13
-rw-r--r--powerpc/SelectOpproof.v26
-rw-r--r--powerpc/TargetPrinter.ml4
-rw-r--r--riscV/Archi.v33
-rw-r--r--riscV/Builtins1.v33
-rw-r--r--riscV/CBuiltins.ml7
-rw-r--r--riscV/SelectOp.vp12
-rw-r--r--riscV/SelectOpproof.v30
-rw-r--r--riscV/TargetPrinter.ml4
-rw-r--r--runtime/include/math.h11
-rw-r--r--runtime/powerpc/i64_dtos.s100
-rw-r--r--runtime/powerpc/i64_dtou.s92
-rw-r--r--runtime/powerpc/i64_sar.s60
-rw-r--r--runtime/powerpc/i64_sdiv.s71
-rw-r--r--runtime/powerpc/i64_shl.s64
-rw-r--r--runtime/powerpc/i64_shr.s65
-rw-r--r--runtime/powerpc/i64_smod.s70
-rw-r--r--runtime/powerpc/i64_smulh.s80
-rw-r--r--runtime/powerpc/i64_stod.s67
-rw-r--r--runtime/powerpc/i64_stof.s68
-rw-r--r--runtime/powerpc/i64_udiv.s54
-rw-r--r--runtime/powerpc/i64_udivmod.s234
-rw-r--r--runtime/powerpc/i64_umod.s47
-rw-r--r--runtime/powerpc/i64_umulh.s65
-rw-r--r--runtime/powerpc/i64_utod.s66
-rw-r--r--runtime/powerpc/i64_utof.s64
-rw-r--r--runtime/powerpc/vararg.s163
-rw-r--r--runtime/powerpc64/i64_dtou.s66
-rw-r--r--runtime/powerpc64/i64_stof.s68
-rw-r--r--runtime/powerpc64/i64_utod.s79
-rw-r--r--runtime/powerpc64/i64_utof.s64
-rw-r--r--runtime/powerpc64/vararg.s163
-rw-r--r--test/monniaux/BearSSL/mk/mkT0.cmd32
-rw-r--r--test/monniaux/builtins/fma.c14
-rw-r--r--test/monniaux/latency/latency.s27
-rw-r--r--test/monniaux/predicated/predicated.s13
-rw-r--r--test/mppa/instr/builtin32.c12
-rw-r--r--test/regression/builtins-arm.c11
-rw-r--r--test/regression/builtins-powerpc.c15
-rw-r--r--test/regression/builtins-riscV.c14
-rw-r--r--test/regression/builtins-x86.c19
-rw-r--r--test/regression/extasm.c10
-rw-r--r--test/regression/ifconv.c20
-rw-r--r--x86/Builtins1.v54
-rw-r--r--x86/CBuiltins.ml6
-rw-r--r--x86/ConstpropOp.vp12
-rw-r--r--x86/ConstpropOpproof.v26
-rw-r--r--x86/SelectOp.vp34
-rw-r--r--x86/SelectOpproof.v42
-rw-r--r--x86/TargetPrinter.ml38
-rw-r--r--x86_32/Archi.v35
-rw-r--r--x86_64/Archi.v35
180 files changed, 6667 insertions, 6367 deletions
diff --git a/LICENSE b/LICENSE
index 5c7d7294..5a7ae79f 100644
--- a/LICENSE
+++ b/LICENSE
@@ -1,6 +1,6 @@
All files in this distribution are part of the CompCert verified compiler.
-The CompCert verified compiler is Copyright by Institut National de
+The CompCert verified compiler is Copyright by Institut National de
Recherche en Informatique et en Automatique (INRIA) and
AbsInt Angewandte Informatik GmbH.
@@ -9,12 +9,12 @@ INRIA Non-Commercial License Agreement given below or under the terms
of a Software Usage Agreement of AbsInt Angewandte Informatik GmbH.
The latter is a separate contract document.
-The INRIA Non-Commercial License Agreement is a non-free license that
-grants you the right to use the CompCert verified compiler for
-educational, research or evaluation purposes only, but prohibits
+The INRIA Non-Commercial License Agreement is a non-free license that
+grants you the right to use the CompCert verified compiler for
+educational, research or evaluation purposes only, but prohibits
any commercial use.
-For commercial use you need a Software Usage Agreement from
+For commercial use you need a Software Usage Agreement from
AbsInt Angewandte Informatik GmbH.
The following files in this distribution are dual-licensed both under
@@ -38,7 +38,7 @@ option) any later version:
cfrontend/Ctyping.v
cfrontend/PrintClight.ml
cfrontend/PrintCsyntax.ml
-
+
backend/Cminor.v
backend/PrintCminor.ml
@@ -46,7 +46,7 @@ option) any later version:
all files in the exportclight/ directory
- the Archi.v, CBuiltins.ml, and extractionMachdep.v files
+ the Archi.v, CBuiltins.ml, and extractionMachdep.v files
in directories arm, powerpc, riscV, x86, x86_32, x86_64
extraction/extraction.v
@@ -64,11 +64,14 @@ non-commercial contexts, subject to the terms of the GNU General
Public License.
The files contained in the flocq/ directory and its subdirectories are
-taken from the Flocq project, http://flocq.gforge.inria.fr/
-These files are Copyright 2010-2017 INRIA and distributed under the
-terms of the GNU Lesser General Public Licence, either version 3 of
-the licence, or (at your option) any later version. A copy of the GNU
-Lesser General Public Licence version 3 is included below.
+taken from the Flocq project, http://flocq.gforge.inria.fr/. The files
+contained in the MenhirLib directory are taken from the Menhir
+project, http://gallium.inria.fr/~fpottier/menhir/. The files from the
+Flocq project and the files in the MenhirLib directory are Copyright
+2010-2019 INRIA and distributed under the terms of the GNU Lesser
+General Public Licence, either version 3 of the licence, or (at your
+option) any later version. A copy of the GNU Lesser General Public
+Licence version 3 is included below.
The files contained in the runtime/ directory and its subdirectories
are Copyright 2013-2017 INRIA and distributed under the terms of the BSD
diff --git a/Makefile b/Makefile
index 104a5546..aa99786b 100644
--- a/Makefile
+++ b/Makefile
@@ -25,12 +25,14 @@ BACKENDLIB?=Asmgenproof0.v Asmgenproof1.v
DIRS=lib common $(ARCHDIRS) backend cfrontend driver \
flocq/Core flocq/Prop flocq/Calc flocq/IEEE754 \
- exportclight cparser cparser/MenhirLib
+ exportclight MenhirLib cparser
-RECDIRS=lib common $(ARCHDIRS) backend cfrontend driver flocq exportclight cparser
+RECDIRS=lib common $(ARCHDIRS) backend cfrontend driver flocq exportclight \
+ MenhirLib cparser
COQINCLUDES=$(foreach d, $(RECDIRS), -R $(d) $(subst /,.,compcert.$(d)))
+COQCOPTS ?= -w -undeclared-scope
COQC="$(COQBIN)coqc" -q $(COQINCLUDES) $(COQCOPTS)
COQDEP="$(COQBIN)coqdep" $(COQINCLUDES)
COQDOC="$(COQBIN)coqdoc"
@@ -64,7 +66,7 @@ VLIB=Axioms.v Coqlib.v Intv.v Maps.v Heaps.v Lattice.v Ordered.v \
COMMON=Errors.v AST.v Linking.v \
Events.v Globalenvs.v Memdata.v Memtype.v Memory.v \
Values.v Smallstep.v Behaviors.v Switch.v Determinism.v Unityping.v \
- Separation.v
+ Separation.v Builtins0.v Builtins1.v Builtins.v
# Back-end modules (in backend/, $(ARCH)/)
@@ -106,16 +108,16 @@ CFRONTEND=Ctypes.v Cop.v Csyntax.v Csem.v Ctyping.v Cstrategy.v Cexec.v \
Cshmgen.v Cshmgenproof.v \
Csharpminor.v Cminorgen.v Cminorgenproof.v
-# LR(1) parser validator
-
-PARSERVALIDATOR=Alphabet.v Interpreter_complete.v Interpreter.v \
- Validator_complete.v Automaton.v Interpreter_correct.v Main.v \
- Validator_safe.v Grammar.v Interpreter_safe.v Tuples.v
-
# Parser
PARSER=Cabs.v Parser.v
+# MenhirLib
+
+MENHIRLIB=Alphabet.v Automaton.v Grammar.v Interpreter_complete.v \
+ Interpreter_correct.v Interpreter.v Main.v Validator_complete.v \
+ Validator_safe.v Validator_classes.v
+
# Putting everything together (in driver/)
DRIVER=Compopts.v Compiler.v Complements.v
@@ -123,7 +125,7 @@ DRIVER=Compopts.v Compiler.v Complements.v
# All source files
FILES=$(VLIB) $(COMMON) $(BACKEND) $(CFRONTEND) $(DRIVER) $(FLOCQ) \
- $(PARSERVALIDATOR) $(PARSER)
+ $(MENHIRLIB) $(PARSER)
# Generated source files
@@ -145,7 +147,6 @@ ifeq ($(CLIGHTGEN),true)
$(MAKE) clightgen
endif
-
proof: $(FILES:.v=.vo)
# Turn off some warnings for compiling Flocq
@@ -229,7 +230,7 @@ driver/Version.ml: VERSION
cparser/Parser.v: cparser/Parser.vy
@rm -f $@
- $(MENHIR) $(MENHIR_FLAGS) --coq cparser/Parser.vy
+ $(MENHIR) --coq --coq-lib-path compcert.MenhirLib --coq-no-version-check cparser/Parser.vy
@chmod a-w $@
depend: $(GENERATED) depend1
diff --git a/Makefile.extr b/Makefile.extr
index d6a94d2e..51e9be59 100644
--- a/Makefile.extr
+++ b/Makefile.extr
@@ -50,8 +50,9 @@ INCLUDES=$(patsubst %,-I %, $(DIRS))
# Control of warnings:
WARNINGS=-w +a-4-9-27-42 -strict-sequence -safe-string -warn-error +a #Deprication returns with ocaml 4.03
-extraction/%.cmx: WARNINGS +=-w -20-27-32..34-39-41-44..45
-extraction/%.cmo: WARNINGS +=-w -20-27-32..34-39-41-44..45
+extraction/%.cmx: WARNINGS +=-w -20-27-32..34-39-41-44..45-60
+extraction/%.cmo: WARNINGS +=-w -20-27-32..34-39-41-44..45-60
+
cparser/pre_parser.cmx: WARNINGS += -w -41
cparser/pre_parser.cmo: WARNINGS += -w -41
diff --git a/MenhirLib/Alphabet.v b/MenhirLib/Alphabet.v
new file mode 100644
index 00000000..29070e3d
--- /dev/null
+++ b/MenhirLib/Alphabet.v
@@ -0,0 +1,247 @@
+(****************************************************************************)
+(* *)
+(* Menhir *)
+(* *)
+(* Jacques-Henri Jourdan, CNRS, LRI, Université Paris Sud *)
+(* *)
+(* Copyright Inria. All rights reserved. This file is distributed under *)
+(* the terms of the GNU Lesser General Public License as published by the *)
+(* Free Software Foundation, either version 3 of the License, or (at your *)
+(* option) any later version, as described in the file LICENSE. *)
+(* *)
+(****************************************************************************)
+
+From Coq Require Import Omega List Syntax Relations RelationClasses.
+
+Local Obligation Tactic := intros.
+
+(** A comparable type is equiped with a [compare] function, that define an order
+ relation. **)
+Class Comparable (A:Type) := {
+ compare : A -> A -> comparison;
+ compare_antisym : forall x y, CompOpp (compare x y) = compare y x;
+ compare_trans : forall x y z c,
+ (compare x y) = c -> (compare y z) = c -> (compare x z) = c
+}.
+
+Theorem compare_refl {A:Type} (C: Comparable A) :
+ forall x, compare x x = Eq.
+Proof.
+intros.
+pose proof (compare_antisym x x).
+destruct (compare x x); intuition; try discriminate.
+Qed.
+
+(** The corresponding order is a strict order. **)
+Definition comparableLt {A:Type} (C: Comparable A) : relation A :=
+ fun x y => compare x y = Lt.
+
+Instance ComparableLtStrictOrder {A:Type} (C: Comparable A) :
+ StrictOrder (comparableLt C).
+Proof.
+apply Build_StrictOrder.
+unfold Irreflexive, Reflexive, complement, comparableLt.
+intros.
+pose proof H.
+rewrite <- compare_antisym in H.
+rewrite H0 in H.
+discriminate H.
+unfold Transitive, comparableLt.
+intros x y z.
+apply compare_trans.
+Qed.
+
+(** nat is comparable. **)
+Program Instance natComparable : Comparable nat :=
+ { compare := Nat.compare }.
+Next Obligation.
+symmetry.
+destruct (Nat.compare x y) as [] eqn:?.
+rewrite Nat.compare_eq_iff in Heqc.
+destruct Heqc.
+rewrite Nat.compare_eq_iff.
+trivial.
+rewrite <- nat_compare_lt in *.
+rewrite <- nat_compare_gt in *.
+trivial.
+rewrite <- nat_compare_lt in *.
+rewrite <- nat_compare_gt in *.
+trivial.
+Qed.
+Next Obligation.
+destruct c.
+rewrite Nat.compare_eq_iff in *; destruct H; assumption.
+rewrite <- nat_compare_lt in *.
+apply (Nat.lt_trans _ _ _ H H0).
+rewrite <- nat_compare_gt in *.
+apply (gt_trans _ _ _ H H0).
+Qed.
+
+(** A pair of comparable is comparable. **)
+Program Instance PairComparable {A:Type} (CA:Comparable A) {B:Type} (CB:Comparable B) :
+ Comparable (A*B) :=
+ { compare := fun x y =>
+ let (xa, xb) := x in let (ya, yb) := y in
+ match compare xa ya return comparison with
+ | Eq => compare xb yb
+ | x => x
+ end }.
+Next Obligation.
+destruct x, y.
+rewrite <- (compare_antisym a a0).
+rewrite <- (compare_antisym b b0).
+destruct (compare a a0); intuition.
+Qed.
+Next Obligation.
+destruct x, y, z.
+destruct (compare a a0) as [] eqn:?, (compare a0 a1) as [] eqn:?;
+try (rewrite <- H0 in H; discriminate);
+try (destruct (compare a a1) as [] eqn:?;
+ try (rewrite <- compare_antisym in Heqc0;
+ rewrite CompOpp_iff in Heqc0;
+ rewrite (compare_trans _ _ _ _ Heqc0 Heqc2) in Heqc1;
+ discriminate);
+ try (rewrite <- compare_antisym in Heqc1;
+ rewrite CompOpp_iff in Heqc1;
+ rewrite (compare_trans _ _ _ _ Heqc2 Heqc1) in Heqc0;
+ discriminate);
+ assumption);
+rewrite (compare_trans _ _ _ _ Heqc0 Heqc1);
+try assumption.
+apply (compare_trans _ _ _ _ H H0).
+Qed.
+
+(** Special case of comparable, where equality is Leibniz equality. **)
+Class ComparableLeibnizEq {A:Type} (C: Comparable A) :=
+ compare_eq : forall x y, compare x y = Eq -> x = y.
+
+(** Boolean equality for a [Comparable]. **)
+Definition compare_eqb {A:Type} {C:Comparable A} (x y:A) :=
+ match compare x y with
+ | Eq => true
+ | _ => false
+ end.
+
+Theorem compare_eqb_iff {A:Type} {C:Comparable A} {U:ComparableLeibnizEq C} :
+ forall x y, compare_eqb x y = true <-> x = y.
+Proof.
+unfold compare_eqb.
+intuition.
+apply compare_eq.
+destruct (compare x y); intuition; discriminate.
+destruct H.
+rewrite compare_refl; intuition.
+Qed.
+
+Instance NComparableLeibnizEq : ComparableLeibnizEq natComparable := Nat.compare_eq.
+
+(** A pair of ComparableLeibnizEq is ComparableLeibnizEq **)
+Instance PairComparableLeibnizEq
+ {A:Type} {CA:Comparable A} (UA:ComparableLeibnizEq CA)
+ {B:Type} {CB:Comparable B} (UB:ComparableLeibnizEq CB) :
+ ComparableLeibnizEq (PairComparable CA CB).
+Proof.
+intros x y; destruct x, y; simpl.
+pose proof (compare_eq a a0); pose proof (compare_eq b b0).
+destruct (compare a a0); try discriminate.
+intuition.
+destruct H2, H0.
+reflexivity.
+Qed.
+
+(** An [Finite] type is a type with the list of all elements. **)
+Class Finite (A:Type) := {
+ all_list : list A;
+ all_list_forall : forall x:A, In x all_list
+}.
+
+(** An alphabet is both [ComparableLeibnizEq] and [Finite]. **)
+Class Alphabet (A:Type) := {
+ AlphabetComparable :> Comparable A;
+ AlphabetComparableLeibnizEq :> ComparableLeibnizEq AlphabetComparable;
+ AlphabetFinite :> Finite A
+}.
+
+(** The [Numbered] class provides a conveniant way to build [Alphabet] instances,
+ with a good computationnal complexity. It is mainly a injection from it to
+ [positive] **)
+Class Numbered (A:Type) := {
+ inj : A -> positive;
+ surj : positive -> A;
+ surj_inj_compat : forall x, surj (inj x) = x;
+ inj_bound : positive;
+ inj_bound_spec : forall x, (inj x < Pos.succ inj_bound)%positive
+}.
+
+Program Instance NumberedAlphabet {A:Type} (N:Numbered A) : Alphabet A :=
+ { AlphabetComparable := {| compare := fun x y => Pos.compare (inj x) (inj y) |};
+ AlphabetFinite :=
+ {| all_list := fst (Pos.iter
+ (fun '(l, p) => (surj p::l, Pos.succ p))
+ ([], 1%positive) inj_bound) |} }.
+Next Obligation. simpl. now rewrite <- Pos.compare_antisym. Qed.
+Next Obligation.
+ match goal with c : comparison |- _ => destruct c end.
+ - rewrite Pos.compare_eq_iff in *. congruence.
+ - rewrite Pos.compare_lt_iff in *. eauto using Pos.lt_trans.
+ - rewrite Pos.compare_gt_iff in *. eauto using Pos.lt_trans.
+Qed.
+Next Obligation.
+ intros x y. unfold compare. intros Hxy.
+ assert (Hxy' : inj x = inj y).
+ (* We do not use [Pos.compare_eq_iff] directly to make sure the
+ proof is executable. *)
+ { destruct (Pos.eq_dec (inj x) (inj y)) as [|[]]; [now auto|].
+ now apply Pos.compare_eq_iff. }
+ (* Using rewrite here leads to non-executable proofs. *)
+ transitivity (surj (inj x)).
+ { apply eq_sym, surj_inj_compat. }
+ transitivity (surj (inj y)); cycle 1.
+ { apply surj_inj_compat. }
+ apply f_equal, Hxy'.
+Defined.
+Next Obligation.
+ rewrite <-(surj_inj_compat x).
+ generalize (inj_bound_spec x). generalize (inj x). clear x. intros x.
+ match goal with |- ?Hx -> In ?s (fst ?p) =>
+ assert ((Hx -> In s (fst p)) /\ snd p = Pos.succ inj_bound); [|now intuition] end.
+ rewrite Pos.lt_succ_r.
+ induction inj_bound as [|y [IH1 IH2]] using Pos.peano_ind;
+ (split; [intros Hx|]); simpl.
+ - rewrite (Pos.le_antisym _ _ Hx); auto using Pos.le_1_l.
+ - auto.
+ - rewrite Pos.iter_succ. destruct Pos.iter; simpl in *. subst.
+ rewrite Pos.le_lteq in Hx. destruct Hx as [?%Pos.lt_succ_r| ->]; now auto.
+ - rewrite Pos.iter_succ. destruct Pos.iter. simpl in IH2. subst. reflexivity.
+Qed.
+
+(** Definitions of [FSet]/[FMap] from [Comparable] **)
+Require Import OrderedTypeAlt.
+Require FSetAVL.
+Require FMapAVL.
+Import OrderedType.
+
+Module Type ComparableM.
+ Parameter t : Type.
+ Declare Instance tComparable : Comparable t.
+End ComparableM.
+
+Module OrderedTypeAlt_from_ComparableM (C:ComparableM) <: OrderedTypeAlt.
+ Definition t := C.t.
+ Definition compare : t -> t -> comparison := compare.
+
+ Infix "?=" := compare (at level 70, no associativity).
+
+ Lemma compare_sym x y : (y?=x) = CompOpp (x?=y).
+ Proof. exact (Logic.eq_sym (compare_antisym x y)). Qed.
+ Lemma compare_trans c x y z :
+ (x?=y) = c -> (y?=z) = c -> (x?=z) = c.
+ Proof.
+ apply compare_trans.
+ Qed.
+End OrderedTypeAlt_from_ComparableM.
+
+Module OrderedType_from_ComparableM (C:ComparableM) <: OrderedType.
+ Module Alt := OrderedTypeAlt_from_ComparableM C.
+ Include (OrderedType_from_Alt Alt).
+End OrderedType_from_ComparableM.
diff --git a/cparser/MenhirLib/Automaton.v b/MenhirLib/Automaton.v
index fc995298..d5a19f35 100644
--- a/cparser/MenhirLib/Automaton.v
+++ b/MenhirLib/Automaton.v
@@ -1,23 +1,20 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Jacques-Henri Jourdan, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
+(****************************************************************************)
+(* *)
+(* Menhir *)
+(* *)
+(* Jacques-Henri Jourdan, CNRS, LRI, Université Paris Sud *)
+(* *)
+(* Copyright Inria. All rights reserved. This file is distributed under *)
+(* the terms of the GNU Lesser General Public License as published by the *)
+(* Free Software Foundation, either version 3 of the License, or (at your *)
+(* option) any later version, as described in the file LICENSE. *)
+(* *)
+(****************************************************************************)
Require Grammar.
-Require Import Orders.
Require Export Alphabet.
-Require Export List.
-Require Export Syntax.
+From Coq Require Import Orders.
+From Coq Require Export List Syntax.
Module Type AutInit.
(** The grammar of the automaton. **)
@@ -102,9 +99,9 @@ Module Types(Import Init:AutInit).
T term = last_symb_of_non_init_state s -> lookahead_action term
| Reduce_act: production -> lookahead_action term
| Fail_act: lookahead_action term.
- Arguments Shift_act [term].
- Arguments Reduce_act [term].
- Arguments Fail_act [term].
+ Arguments Shift_act {term}.
+ Arguments Reduce_act {term}.
+ Arguments Fail_act {term}.
Inductive action :=
| Default_reduce_act: production -> action
diff --git a/MenhirLib/Grammar.v b/MenhirLib/Grammar.v
new file mode 100644
index 00000000..a371318b
--- /dev/null
+++ b/MenhirLib/Grammar.v
@@ -0,0 +1,162 @@
+(****************************************************************************)
+(* *)
+(* Menhir *)
+(* *)
+(* Jacques-Henri Jourdan, CNRS, LRI, Université Paris Sud *)
+(* *)
+(* Copyright Inria. All rights reserved. This file is distributed under *)
+(* the terms of the GNU Lesser General Public License as published by the *)
+(* Free Software Foundation, either version 3 of the License, or (at your *)
+(* option) any later version, as described in the file LICENSE. *)
+(* *)
+(****************************************************************************)
+
+From Coq Require Import List Syntax Orders.
+Require Import Alphabet.
+
+(** The terminal non-terminal alphabets of the grammar. **)
+Module Type Alphs.
+ Parameters terminal nonterminal : Type.
+ Declare Instance TerminalAlph: Alphabet terminal.
+ Declare Instance NonTerminalAlph: Alphabet nonterminal.
+End Alphs.
+
+(** Definition of the alphabet of symbols, given the alphabet of terminals
+ and the alphabet of non terminals **)
+Module Symbol(Import A:Alphs).
+
+ Inductive symbol :=
+ | T: terminal -> symbol
+ | NT: nonterminal -> symbol.
+
+ Program Instance SymbolAlph : Alphabet symbol :=
+ { AlphabetComparable := {| compare := fun x y =>
+ match x, y return comparison with
+ | T _, NT _ => Gt
+ | NT _, T _ => Lt
+ | T x, T y => compare x y
+ | NT x, NT y => compare x y
+ end |};
+ AlphabetFinite := {| all_list :=
+ map T all_list++map NT all_list |} }.
+ Next Obligation.
+ destruct x; destruct y; intuition; apply compare_antisym.
+ Qed.
+ Next Obligation.
+ destruct x; destruct y; destruct z; intuition; try discriminate.
+ apply (compare_trans _ t0); intuition.
+ apply (compare_trans _ n0); intuition.
+ Qed.
+ Next Obligation.
+ intros x y.
+ destruct x; destruct y; try discriminate; intros.
+ rewrite (compare_eq t t0); now intuition.
+ rewrite (compare_eq n n0); now intuition.
+ Defined.
+ Next Obligation.
+ rewrite in_app_iff.
+ destruct x; [left | right]; apply in_map; apply all_list_forall.
+ Qed.
+
+End Symbol.
+
+(** A curryfied function with multiple parameters **)
+Definition arrows_right: Type -> list Type -> Type :=
+ fold_right (fun A B => A -> B).
+
+Module Type T.
+ Include Alphs <+ Symbol.
+
+ (** [symbol_semantic_type] maps a symbols to the type of its semantic
+ values. **)
+ Parameter symbol_semantic_type: symbol -> Type.
+
+ (** The type of productions identifiers **)
+ Parameter production : Type.
+ Declare Instance ProductionAlph : Alphabet production.
+
+ (** Accessors for productions: left hand side, right hand side,
+ and semantic action. The semantic actions are given in the form
+ of curryfied functions, that take arguments in the reverse order. **)
+ Parameter prod_lhs: production -> nonterminal.
+ (* The RHS of a production is given in reversed order, so that symbols *)
+ Parameter prod_rhs_rev: production -> list symbol.
+ Parameter prod_action:
+ forall p:production,
+ arrows_right
+ (symbol_semantic_type (NT (prod_lhs p)))
+ (map symbol_semantic_type (prod_rhs_rev p)).
+
+ (** Tokens are the atomic elements of the input stream: they contain
+ a terminal and a semantic value of the type corresponding to this
+ terminal. *)
+ Parameter token : Type.
+ Parameter token_term : token -> terminal.
+ Parameter token_sem :
+ forall tok : token, symbol_semantic_type (T (token_term tok)).
+End T.
+
+Module Defs(Import G:T).
+
+ (** The semantics of a grammar is defined in two stages. First, we
+ define the notion of parse tree, which represents one way of
+ recognizing a word with a head symbol. Semantic values are stored
+ at the leaves.
+
+ This notion is defined in two mutually recursive flavours:
+ either for a single head symbol, or for a list of head symbols. *)
+ Inductive parse_tree:
+ forall (head_symbol:symbol) (word:list token), Type :=
+
+ (** Parse tree for a terminal symbol. *)
+ | Terminal_pt:
+ forall (tok:token), parse_tree (T (token_term tok)) [tok]
+
+ (** Parse tree for a non-terminal symbol. *)
+ | Non_terminal_pt:
+ forall (prod:production) {word:list token},
+ parse_tree_list (prod_rhs_rev prod) word ->
+ parse_tree (NT (prod_lhs prod)) word
+
+ (* Note : the list head_symbols_rev is reversed. *)
+ with parse_tree_list:
+ forall (head_symbols_rev:list symbol) (word:list token), Type :=
+
+ | Nil_ptl: parse_tree_list [] []
+
+ | Cons_ptl:
+ forall {head_symbolsq:list symbol} {wordq:list token},
+ parse_tree_list head_symbolsq wordq ->
+
+ forall {head_symbolt:symbol} {wordt:list token},
+ parse_tree head_symbolt wordt ->
+
+ parse_tree_list (head_symbolt::head_symbolsq) (wordq++wordt).
+
+ (** We can now finish the definition of the semantics of a grammar,
+ by giving the semantic value assotiated with a parse tree. *)
+ Fixpoint pt_sem {head_symbol word} (tree:parse_tree head_symbol word) :
+ symbol_semantic_type head_symbol :=
+ match tree with
+ | Terminal_pt tok => token_sem tok
+ | Non_terminal_pt prod ptl => ptl_sem ptl (prod_action prod)
+ end
+ with ptl_sem {A head_symbols word} (tree:parse_tree_list head_symbols word) :
+ arrows_right A (map symbol_semantic_type head_symbols) -> A :=
+ match tree with
+ | Nil_ptl => fun act => act
+ | Cons_ptl q t => fun act => ptl_sem q (act (pt_sem t))
+ end.
+
+ Fixpoint pt_size {head_symbol word} (tree:parse_tree head_symbol word) :=
+ match tree with
+ | Terminal_pt _ => 1
+ | Non_terminal_pt _ l => S (ptl_size l)
+ end
+ with ptl_size {head_symbols word} (tree:parse_tree_list head_symbols word) :=
+ match tree with
+ | Nil_ptl => 0
+ | Cons_ptl q t =>
+ pt_size t + ptl_size q
+ end.
+End Defs.
diff --git a/MenhirLib/Interpreter.v b/MenhirLib/Interpreter.v
new file mode 100644
index 00000000..568597ba
--- /dev/null
+++ b/MenhirLib/Interpreter.v
@@ -0,0 +1,453 @@
+(****************************************************************************)
+(* *)
+(* Menhir *)
+(* *)
+(* Jacques-Henri Jourdan, CNRS, LRI, Université Paris Sud *)
+(* *)
+(* Copyright Inria. All rights reserved. This file is distributed under *)
+(* the terms of the GNU Lesser General Public License as published by the *)
+(* Free Software Foundation, either version 3 of the License, or (at your *)
+(* option) any later version, as described in the file LICENSE. *)
+(* *)
+(****************************************************************************)
+
+From Coq Require Import List Syntax.
+From Coq.ssr Require Import ssreflect.
+Require Automaton.
+Require Import Alphabet Grammar Validator_safe.
+
+Module Make(Import A:Automaton.T).
+Module Import ValidSafe := Validator_safe.Make A.
+
+(** A few helpers for dependent types. *)
+
+(** Decidable propositions. *)
+Class Decidable (P : Prop) := decide : {P} + {~P}.
+Arguments decide _ {_}.
+
+(** A [Comparable] type has decidable equality. *)
+Instance comparable_decidable_eq T `{ComparableLeibnizEq T} (x y : T) :
+ Decidable (x = y).
+Proof.
+ unfold Decidable.
+ destruct (compare x y) eqn:EQ; [left; apply compare_eq; intuition | ..];
+ right; intros ->; by rewrite compare_refl in EQ.
+Defined.
+
+Instance list_decidable_eq T :
+ (forall x y : T, Decidable (x = y)) ->
+ (forall l1 l2 : list T, Decidable (l1 = l2)).
+Proof. unfold Decidable. decide equality. Defined.
+
+Ltac subst_existT :=
+ repeat
+ match goal with
+ | _ => progress subst
+ | H : @existT ?A ?P ?x ?y1 = @existT ?A ?P ?x ?y2 |- _ =>
+ let DEC := fresh in
+ assert (DEC : forall u1 u2 : A, Decidable (u1 = u2)) by apply _;
+ apply Eqdep_dec.inj_pair2_eq_dec in H; [|by apply DEC];
+ clear DEC
+ end.
+
+(** The interpreter is written using dependent types. In order to
+ avoid reducing proof terms while executing the parser, we thunk all
+ the propositions behind an arrow.
+ Note that thunkP is still in Prop so that it is erased by
+ extraction.
+ *)
+Definition thunkP (P : Prop) : Prop := True -> P.
+
+(** Sometimes, we actually need a reduced proof in a program (for
+ example when using an equality to cast a value). In that case,
+ instead of reducing the proof we already have, we reprove the
+ assertion by using decidability. *)
+Definition reprove {P} `{Decidable P} (p : thunkP P) : P :=
+ match decide P with
+ | left p => p
+ | right np => False_ind _ (np (p I))
+ end.
+
+(** Combination of reprove with eq_rect. *)
+Definition cast {T : Type} (F : T -> Type) {x y : T} (eq : thunkP (x = y))
+ {DEC : unit -> Decidable (x = y)}:
+ F x -> F y :=
+ fun a => eq_rect x F a y (@reprove _ (DEC ()) eq).
+
+Lemma cast_eq T F (x : T) (eq : thunkP (x = x)) `{forall x y, Decidable (x = y)} a :
+ cast F eq a = a.
+Proof. by rewrite /cast -Eqdep_dec.eq_rect_eq_dec. Qed.
+
+(** Input buffers and operations on them. **)
+CoInductive buffer : Type :=
+ Buf_cons { buf_head : token; buf_tail : buffer }.
+
+Delimit Scope buffer_scope with buf.
+Bind Scope buffer_scope with buffer.
+
+Infix "::" := Buf_cons (at level 60, right associativity) : buffer_scope.
+
+(** Concatenation of a list and an input buffer **)
+Fixpoint app_buf (l:list token) (buf:buffer) :=
+ match l with
+ | nil => buf
+ | cons t q => (t :: app_buf q buf)%buf
+ end.
+Infix "++" := app_buf (at level 60, right associativity) : buffer_scope.
+
+Lemma app_buf_assoc (l1 l2:list token) (buf:buffer) :
+ (l1 ++ (l2 ++ buf) = (l1 ++ l2) ++ buf)%buf.
+Proof. induction l1 as [|?? IH]=>//=. rewrite IH //. Qed.
+
+(** The type of a non initial state: the type of semantic values associated
+ with the last symbol of this state. *)
+Definition noninitstate_type state :=
+ symbol_semantic_type (last_symb_of_non_init_state state).
+
+(** The stack of the automaton : it can be either nil or contains a non
+ initial state, a semantic value for the symbol associted with this state,
+ and a nested stack. **)
+Definition stack := list (sigT noninitstate_type). (* eg. list {state & state_type state} *)
+
+Section Interpreter.
+
+Hypothesis safe: safe.
+
+(* Properties of the automaton deduced from safety validation. *)
+Proposition shift_head_symbs: shift_head_symbs.
+Proof. pose proof safe; unfold ValidSafe.safe in H; intuition. Qed.
+Proposition goto_head_symbs: goto_head_symbs.
+Proof. pose proof safe; unfold ValidSafe.safe in H; intuition. Qed.
+Proposition shift_past_state: shift_past_state.
+Proof. pose proof safe; unfold ValidSafe.safe in H; intuition. Qed.
+Proposition goto_past_state: goto_past_state.
+Proof. pose proof safe; unfold ValidSafe.safe in H; intuition. Qed.
+Proposition reduce_ok: reduce_ok.
+Proof. pose proof safe; unfold ValidSafe.safe in H; intuition. Qed.
+
+Variable init : initstate.
+
+(** The top state of a stack **)
+Definition state_of_stack (stack:stack): state :=
+ match stack with
+ | [] => init
+ | existT _ s _::_ => s
+ end.
+
+(** The stack of states of an automaton stack **)
+Definition state_stack_of_stack (stack:stack) :=
+ (List.map
+ (fun cell:sigT noninitstate_type => singleton_state_pred (projT1 cell))
+ stack ++ [singleton_state_pred init])%list.
+
+(** The stack of symbols of an automaton stack **)
+Definition symb_stack_of_stack (stack:stack) :=
+ List.map
+ (fun cell:sigT noninitstate_type => last_symb_of_non_init_state (projT1 cell))
+ stack.
+
+(** The stack invariant : it basically states that the assumptions on the
+ states are true. **)
+Inductive stack_invariant: stack -> Prop :=
+ | stack_invariant_constr:
+ forall stack,
+ prefix (head_symbs_of_state (state_of_stack stack))
+ (symb_stack_of_stack stack) ->
+ prefix_pred (head_states_of_state (state_of_stack stack))
+ (state_stack_of_stack stack) ->
+ stack_invariant_next stack ->
+ stack_invariant stack
+with stack_invariant_next: stack -> Prop :=
+ | stack_invariant_next_nil:
+ stack_invariant_next []
+ | stack_invariant_next_cons:
+ forall state_cur st stack_rec,
+ stack_invariant stack_rec ->
+ stack_invariant_next (existT _ state_cur st::stack_rec).
+
+(** [pop] pops some symbols from the stack. It returns the popped semantic
+ values using [sem_popped] as an accumulator and discards the popped
+ states.**)
+Fixpoint pop (symbols_to_pop:list symbol) {A:Type} (stk:stack) :
+ thunkP (prefix symbols_to_pop (symb_stack_of_stack stk)) ->
+ forall (action:arrows_right A (map symbol_semantic_type symbols_to_pop)),
+ stack * A.
+unshelve refine
+ (match symbols_to_pop
+ return
+ (thunkP (prefix symbols_to_pop (symb_stack_of_stack stk))) ->
+ forall (action:arrows_right A (map _ symbols_to_pop)), stack * A
+ with
+ | [] => fun _ action => (stk, action)
+ | t::q => fun Hp action =>
+ match stk
+ return thunkP (prefix (t::q) (symb_stack_of_stack stk)) -> stack * A
+ with
+ | existT _ state_cur sem::stack_rec => fun Hp =>
+ let sem_conv := cast symbol_semantic_type _ sem in
+ pop q _ stack_rec _ (action sem_conv)
+ | [] => fun Hp => False_rect _ _
+ end Hp
+ end).
+Proof.
+ - simpl in Hp. clear -Hp. abstract (intros _ ; specialize (Hp I); now inversion Hp).
+ - clear -Hp. abstract (specialize (Hp I); now inversion Hp).
+ - simpl in Hp. clear -Hp. abstract (intros _ ; specialize (Hp I); now inversion Hp).
+Defined.
+
+(* Equivalent declarative specification for pop, so that we avoid
+ (part of) the dependent types nightmare. *)
+Inductive pop_spec {A:Type} :
+ forall (symbols_to_pop:list symbol) (stk : stack)
+ (action : arrows_right A (map symbol_semantic_type symbols_to_pop))
+ (stk' : stack) (sem : A),
+ Prop :=
+ | Nil_pop_spec stk sem : pop_spec [] stk sem stk sem
+ | Cons_pop_spec symbols_to_pop st stk action sem stk' res :
+ pop_spec symbols_to_pop stk (action sem) stk' res ->
+ pop_spec (last_symb_of_non_init_state st::symbols_to_pop)
+ (existT _ st sem :: stk) action stk' res.
+
+Lemma pop_spec_ok {A:Type} symbols_to_pop stk Hp action stk' res:
+ pop symbols_to_pop stk Hp action = (stk', res) <->
+ pop_spec (A:=A) symbols_to_pop stk action stk' res.
+Proof.
+ revert stk Hp action.
+ induction symbols_to_pop as [|t symbols_to_pop IH]=>stk Hp action /=.
+ - split.
+ + intros [= <- <-]. constructor.
+ + intros H. inversion H. by subst_existT.
+ - destruct stk as [|[st sem]]=>/=; [by destruct pop_subproof0|].
+ remember (pop_subproof t symbols_to_pop stk st Hp) as EQ eqn:eq. clear eq.
+ generalize EQ. revert Hp action. rewrite <-(EQ I)=>Hp action ?.
+ rewrite cast_eq. rewrite IH. split.
+ + intros. by constructor.
+ + intros H. inversion H. by subst_existT.
+Qed.
+
+
+Lemma pop_preserves_invariant symbols_to_pop stk Hp A action :
+ stack_invariant stk ->
+ stack_invariant (fst (pop symbols_to_pop stk Hp (A:=A) action)).
+Proof.
+ revert stk Hp A action. induction symbols_to_pop as [|t q IH]=>//=.
+ intros stk Hp A action Hi.
+ destruct Hi as [stack Hp' Hpp [|state st stk']].
+ - destruct pop_subproof0.
+ - now apply IH.
+Qed.
+
+Lemma pop_state_valid symbols_to_pop stk Hp A action lpred :
+ prefix_pred lpred (state_stack_of_stack stk) ->
+ let stk' := fst (pop symbols_to_pop stk Hp (A:=A) action) in
+ state_valid_after_pop (state_of_stack stk') symbols_to_pop lpred.
+Proof.
+ revert stk Hp A action lpred. induction symbols_to_pop as [|t q IH]=>/=.
+ - intros stk Hp A a lpred Hpp. destruct lpred as [|pred lpred]; constructor.
+ inversion Hpp as [|? lpred' ? pred' Himpl Hpp' eq1 eq2]; subst.
+ specialize (Himpl (state_of_stack stk)).
+ destruct (pred' (state_of_stack stk)) as [] eqn:Heqpred'=>//.
+ destruct stk as [|[]]; simpl in *.
+ + inversion eq2; subst; clear eq2.
+ unfold singleton_state_pred in Heqpred'.
+ now rewrite compare_refl in Heqpred'; discriminate.
+ + inversion eq2; subst; clear eq2.
+ unfold singleton_state_pred in Heqpred'.
+ now rewrite compare_refl in Heqpred'; discriminate.
+ - intros stk Hp A a lpred Hpp. destruct stk as [|[] stk]=>//=.
+ + destruct pop_subproof0.
+ + destruct lpred as [|pred lpred]; [by constructor|].
+ constructor. apply IH. by inversion Hpp.
+Qed.
+
+(** [step_result] represents the result of one step of the automaton : it can
+ fail, accept or progress. [Fail_sr] means that the input is incorrect.
+ [Accept_sr] means that this is the last step of the automaton, and it
+ returns the semantic value of the input word. [Progress_sr] means that
+ some progress has been made, but new steps are needed in order to accept
+ a word.
+
+ For [Accept_sr] and [Progress_sr], the result contains the new input buffer.
+
+ [Fail_sr] means that the input word is rejected by the automaton. It is
+ different to [Err] (from the error monad), which mean that the automaton is
+ bogus and has perfomed a forbidden action. **)
+Inductive step_result :=
+ | Fail_sr: step_result
+ | Accept_sr: symbol_semantic_type (NT (start_nt init)) -> buffer -> step_result
+ | Progress_sr: stack -> buffer -> step_result.
+
+(** [reduce_step] does a reduce action :
+ - pops some elements from the stack
+ - execute the action of the production
+ - follows the goto for the produced non terminal symbol **)
+Definition reduce_step stk prod (buffer : buffer)
+ (Hval : thunkP (valid_for_reduce (state_of_stack stk) prod))
+ (Hi : thunkP (stack_invariant stk))
+ : step_result.
+refine
+ ((let '(stk', sem) as ss := pop (prod_rhs_rev prod) stk _ (prod_action prod)
+ return thunkP (state_valid_after_pop (state_of_stack (fst ss)) _
+ (head_states_of_state (state_of_stack stk))) -> _
+ in fun Hval' =>
+ match goto_table (state_of_stack stk') (prod_lhs prod) as goto
+ return (thunkP (goto = None ->
+ match state_of_stack stk' with
+ | Init i => prod_lhs prod = start_nt i
+ | Ninit _ => False
+ end)) -> _
+ with
+ | Some (exist _ state_new e) => fun _ =>
+ let sem := eq_rect _ _ sem _ e in
+ Progress_sr (existT noninitstate_type state_new sem::stk') buffer
+ | None => fun Hval =>
+ let sem := cast symbol_semantic_type _ sem in
+ Accept_sr sem buffer
+ end (fun _ => _))
+ (fun _ => pop_state_valid _ _ _ _ _ _ _)).
+Proof.
+ - clear -Hi Hval.
+ abstract (intros _; destruct Hi=>//; eapply prefix_trans; [by apply Hval|eassumption]).
+ - clear -Hval.
+ abstract (intros _; f_equal; specialize (Hval I eq_refl); destruct stk' as [|[]]=>//).
+ - simpl in Hval'. clear -Hval Hval'.
+ abstract (move : Hval => /(_ I) [_ /(_ _ (Hval' I))] Hval2 Hgoto; by rewrite Hgoto in Hval2).
+ - clear -Hi. abstract by destruct Hi.
+Defined.
+
+Lemma reduce_step_stack_invariant_preserved stk prod buffer Hv Hi stk' buffer':
+ reduce_step stk prod buffer Hv Hi = Progress_sr stk' buffer' ->
+ stack_invariant stk'.
+Proof.
+ unfold reduce_step.
+ match goal with
+ | |- context [pop ?symbols_to_pop stk ?Hp ?action] =>
+ assert (Hi':=pop_preserves_invariant symbols_to_pop stk Hp _ action (Hi I));
+ generalize (pop_state_valid symbols_to_pop stk Hp _ action)
+ end.
+ destruct pop as [stk0 sem]=>/=. simpl in Hi'. intros Hv'.
+ assert (Hgoto1:=goto_head_symbs (state_of_stack stk0) (prod_lhs prod)).
+ assert (Hgoto2:=goto_past_state (state_of_stack stk0) (prod_lhs prod)).
+ match goal with | |- context [fun _ : True => ?X] => generalize X end.
+ destruct goto_table as [[state_new e]|] eqn:EQgoto=>//.
+ intros _ [= <- <-]. constructor=>/=.
+ - constructor. eapply prefix_trans. apply Hgoto1. by destruct Hi'.
+ - unfold state_stack_of_stack; simpl; constructor.
+ + intros ?. by destruct singleton_state_pred.
+ + eapply prefix_pred_trans. apply Hgoto2. by destruct Hi'.
+ - by constructor.
+Qed.
+
+(** One step of parsing. **)
+Definition step stk buffer (Hi : thunkP (stack_invariant stk)): step_result :=
+ match action_table (state_of_stack stk) as a return
+ thunkP
+ match a return Prop with
+ | Default_reduce_act prod => _
+ | Lookahead_act awt => forall t : terminal,
+ match awt t with
+ | Reduce_act p => _
+ | _ => True
+ end
+ end -> _
+ with
+ | Default_reduce_act prod => fun Hv =>
+ reduce_step stk prod buffer Hv Hi
+ | Lookahead_act awt => fun Hv =>
+ match buf_head buffer with
+ | tok =>
+ match awt (token_term tok) as a return
+ thunkP match a return Prop with Reduce_act p => _ | _ => _ end -> _
+ with
+ | Shift_act state_new e => fun _ =>
+ let sem_conv := eq_rect _ symbol_semantic_type (token_sem tok) _ e in
+ Progress_sr (existT noninitstate_type state_new sem_conv::stk)
+ (buf_tail buffer)
+ | Reduce_act prod => fun Hv =>
+ reduce_step stk prod buffer Hv Hi
+ | Fail_act => fun _ =>
+ Fail_sr
+ end (fun _ => Hv I (token_term tok))
+ end
+ end (fun _ => reduce_ok _).
+
+Lemma step_stack_invariant_preserved stk buffer Hi stk' buffer':
+ step stk buffer Hi = Progress_sr stk' buffer' ->
+ stack_invariant stk'.
+Proof.
+ unfold step.
+ generalize (reduce_ok (state_of_stack stk))=>Hred.
+ assert (Hshift1 := shift_head_symbs (state_of_stack stk)).
+ assert (Hshift2 := shift_past_state (state_of_stack stk)).
+ destruct action_table as [prod|awt]=>/=.
+ - eauto using reduce_step_stack_invariant_preserved.
+ - set (term := token_term (buf_head buffer)).
+ generalize (Hred term). clear Hred. intros Hred.
+ specialize (Hshift1 term). specialize (Hshift2 term).
+ destruct (awt term) as [state_new e|prod|]=>//.
+ + intros [= <- <-]. constructor=>/=.
+ * constructor. eapply prefix_trans. apply Hshift1. by destruct Hi.
+ * unfold state_stack_of_stack; simpl; constructor.
+ -- intros ?. by destruct singleton_state_pred.
+ -- eapply prefix_pred_trans. apply Hshift2. by destruct Hi.
+ * constructor; by apply Hi.
+ + eauto using reduce_step_stack_invariant_preserved.
+Qed.
+
+(** The parsing use a [nat] fuel parameter [log_n_steps], so that we
+ do not have to prove terminaison, which is difficult.
+
+ Note that [log_n_steps] is *not* the fuel in the conventionnal
+ sense: this parameter contains the logarithm (in base 2) of the
+ number of steps to perform. Hence, a value of, e.g., 50 will
+ usually be enough to ensure termination. *)
+Fixpoint parse_fix stk buffer (log_n_steps : nat) (Hi : thunkP (stack_invariant stk)):
+ { sr : step_result |
+ forall stk' buffer', sr = Progress_sr stk' buffer' -> stack_invariant stk' } :=
+ match log_n_steps with
+ | O => exist _ (step stk buffer Hi)
+ (step_stack_invariant_preserved _ _ Hi)
+ | S log_n_steps =>
+ match parse_fix stk buffer log_n_steps Hi with
+ | exist _ (Progress_sr stk buffer) Hi' =>
+ parse_fix stk buffer log_n_steps (fun _ => Hi' _ buffer eq_refl)
+ | sr => sr
+ end
+ end.
+
+(** The final result of a parsing is either a failure (the automaton
+ has rejected the input word), either a timeout (the automaton has
+ spent all the given [2^log_n_steps]), either a parsed semantic value
+ with a rest of the input buffer.
+
+ Note that we do not make parse_result depend on start_nt for the
+ result type, so that this inductive is extracted without the use
+ of Obj.t in OCaml. **)
+Inductive parse_result {A : Type} :=
+ | Fail_pr: parse_result
+ | Timeout_pr: parse_result
+ | Parsed_pr: A -> buffer -> parse_result.
+Global Arguments parse_result _ : clear implicits.
+
+Definition parse (buffer : buffer) (log_n_steps : nat):
+ parse_result (symbol_semantic_type (NT (start_nt init))).
+refine (match proj1_sig (parse_fix [] buffer log_n_steps _) with
+ | Fail_sr => Fail_pr
+ | Accept_sr sem buffer' => Parsed_pr sem buffer'
+ | Progress_sr _ _ => Timeout_pr
+ end).
+Proof.
+ abstract (repeat constructor; intros; by destruct singleton_state_pred).
+Defined.
+
+End Interpreter.
+
+Arguments Fail_sr {init}.
+Arguments Accept_sr {init} _ _.
+Arguments Progress_sr {init} _ _.
+
+End Make.
+
+Module Type T(A:Automaton.T).
+ Include (Make A).
+End T.
diff --git a/MenhirLib/Interpreter_complete.v b/MenhirLib/Interpreter_complete.v
new file mode 100644
index 00000000..ec69592b
--- /dev/null
+++ b/MenhirLib/Interpreter_complete.v
@@ -0,0 +1,825 @@
+(****************************************************************************)
+(* *)
+(* Menhir *)
+(* *)
+(* Jacques-Henri Jourdan, CNRS, LRI, Université Paris Sud *)
+(* *)
+(* Copyright Inria. All rights reserved. This file is distributed under *)
+(* the terms of the GNU Lesser General Public License as published by the *)
+(* Free Software Foundation, either version 3 of the License, or (at your *)
+(* option) any later version, as described in the file LICENSE. *)
+(* *)
+(****************************************************************************)
+
+From Coq Require Import List Syntax Arith.
+From Coq.ssr Require Import ssreflect.
+Require Import Alphabet Grammar.
+Require Automaton Interpreter Validator_complete.
+
+Module Make(Import A:Automaton.T) (Import Inter:Interpreter.T A).
+Module Import Valid := Validator_complete.Make A.
+
+(** * Completeness Proof **)
+
+Section Completeness_Proof.
+
+Hypothesis safe: Inter.ValidSafe.safe.
+Hypothesis complete: complete.
+
+(* Properties of the automaton deduced from completeness validation. *)
+Proposition nullable_stable: nullable_stable.
+Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
+Proposition first_stable: first_stable.
+Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
+Proposition start_future: start_future.
+Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
+Proposition terminal_shift: terminal_shift.
+Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
+Proposition end_reduce: end_reduce.
+Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
+Proposition start_goto: start_goto.
+Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
+Proposition non_terminal_goto: non_terminal_goto.
+Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
+Proposition non_terminal_closed: non_terminal_closed.
+Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
+
+(** If the nullable predicate has been validated, then it is correct. **)
+Lemma nullable_correct head word :
+ word = [] -> parse_tree head word -> nullable_symb head = true
+with nullable_correct_list heads word :
+ word = [] ->
+ parse_tree_list heads word -> nullable_word heads = true.
+Proof.
+ - destruct 2=>//. assert (Hnull := nullable_stable prod).
+ erewrite nullable_correct_list in Hnull; eauto.
+ - intros Hword. destruct 1=>//=. destruct (app_eq_nil _ _ Hword).
+ eauto using andb_true_intro.
+Qed.
+
+(** Auxiliary lemma for first_correct. *)
+Lemma first_word_set_app t word1 word2 :
+ TerminalSet.In t (first_word_set (word1 ++ word2)) <->
+ TerminalSet.In t (first_word_set word1) \/
+ TerminalSet.In t (first_word_set word2) /\ nullable_word (rev word1) = true.
+Proof.
+ induction word1 as [|s word1 IH]=>/=.
+ - split; [tauto|]. move=>[/TerminalSet.empty_1 ?|[? _]]//.
+ - rewrite /nullable_word forallb_app /=. destruct nullable_symb=>/=.
+ + rewrite Bool.andb_true_r. split.
+ * move=>/TerminalSet.union_1. rewrite IH.
+ move=>[?|[?|[??]]]; auto using TerminalSet.union_2, TerminalSet.union_3.
+ * destruct IH.
+ move=>[/TerminalSet.union_1 [?|?]|[??]];
+ auto using TerminalSet.union_2, TerminalSet.union_3.
+ + rewrite Bool.andb_false_r. by intuition.
+Qed.
+
+(** If the first predicate has been validated, then it is correct. **)
+Lemma first_correct head word t q :
+ word = t::q ->
+ parse_tree head word ->
+ TerminalSet.In (token_term t) (first_symb_set head)
+with first_correct_list heads word t q :
+ word = t::q ->
+ parse_tree_list heads word ->
+ TerminalSet.In (token_term t) (first_word_set (rev' heads)).
+Proof.
+ - intros Hword. destruct 1=>//.
+ + inversion Hword. subst. apply TerminalSet.singleton_2, compare_refl.
+ + eapply first_stable. eauto.
+ - intros Hword. destruct 1 as [|symq wordq ptl symt wordt pt]=>//=.
+ rewrite /rev' -rev_alt /= first_word_set_app /= rev_involutive rev_alt.
+ destruct wordq; [right|left].
+ + destruct nullable_symb; eauto using TerminalSet.union_2, nullable_correct_list.
+ + inversion Hword. subst. fold (rev' symq). eauto.
+Qed.
+
+(** A PTL is compatible with a stack if the top of the stack contains
+ data representing to this PTL. *)
+Fixpoint ptl_stack_compat {symbs word}
+ (stk0 : stack) (ptl : parse_tree_list symbs word) (stk : stack) : Prop :=
+ match ptl with
+ | Nil_ptl => stk0 = stk
+ | @Cons_ptl _ _ ptl sym _ pt =>
+ match stk with
+ | [] => False
+ | existT _ _ sem::stk =>
+ ptl_stack_compat stk0 ptl stk /\
+ exists e,
+ sem = eq_rect _ symbol_semantic_type (pt_sem pt) _ e
+ end
+ end.
+
+(** .. and when a PTL is compatible with a stack, then calling the pop
+ function return the semantic value of this PTL. *)
+Lemma pop_stack_compat_pop_spec {A symbs word}
+ (ptl:parse_tree_list symbs word) (stk:stack) (stk0:stack) action :
+ ptl_stack_compat stk0 ptl stk ->
+ pop_spec symbs stk action stk0 (ptl_sem (A:=A) ptl action).
+Proof.
+ revert stk. induction ptl=>stk /= Hstk.
+ - subst. constructor.
+ - destruct stk as [|[st sem] stk]=>//. destruct Hstk as [Hstk [??]]. subst.
+ simpl. constructor. eauto.
+Qed.
+
+Variable init: initstate.
+
+(** In order to prove compleness, we first fix a word to be parsed
+ together with the content of the parser at the end of the parsing. *)
+Variable full_word: list token.
+Variable buffer_end: buffer.
+
+(** Completeness is proved by following the traversal of the parse
+ tree which is performed by the parser. Each step of parsing
+ correspond to one step of traversal. In order to represent the state
+ of the traversal, we define the notion of "dotted" parse tree, which
+ is a parse tree with one dot on one of its node. The place of the
+ dot represents the place of the next action to be executed.
+
+ Such a dotted parse tree is decomposed into two part: a "regular"
+ parse tree, which is the parse tree placed under the dot, and a
+ "parse tree zipper", which is the part of the parse tree placed
+ above the dot. Therefore, a parse tree zipper is a parse tree with a
+ hole. Moreover, for easier manipulation, a parse tree zipper is
+ represented "upside down". That is, the root of the parse tree is
+ actually a leaf of the zipper, while the root of the zipper is the
+ hole.
+ *)
+Inductive pt_zipper:
+ forall (hole_symb:symbol) (hole_word:list token), Type :=
+| Top_ptz:
+ pt_zipper (NT (start_nt init)) full_word
+| Cons_ptl_ptz:
+ forall {head_symbolsq:list symbol} {wordq:list token},
+ parse_tree_list head_symbolsq wordq ->
+
+ forall {head_symbolt:symbol} {wordt:list token},
+
+ ptl_zipper (head_symbolt::head_symbolsq) (wordq++wordt) ->
+ pt_zipper head_symbolt wordt
+with ptl_zipper:
+ forall (hole_symbs:list symbol) (hole_word:list token), Type :=
+| Non_terminal_pt_ptlz:
+ forall {p:production} {word:list token},
+ pt_zipper (NT (prod_lhs p)) word ->
+ ptl_zipper (prod_rhs_rev p) word
+
+| Cons_ptl_ptlz:
+ forall {head_symbolsq:list symbol} {wordq:list token},
+
+ forall {head_symbolt:symbol} {wordt:list token},
+ parse_tree head_symbolt wordt ->
+
+ ptl_zipper (head_symbolt::head_symbolsq) (wordq++wordt) ->
+
+ ptl_zipper head_symbolsq wordq.
+
+(** A dotted parse tree is the combination of a parse tree zipper with
+ a parse tree. It can be intwo flavors, depending on which is the next
+ action to be executed (shift or reduce). *)
+Inductive pt_dot: Type :=
+| Reduce_ptd: forall {prod word},
+ parse_tree_list (prod_rhs_rev prod) word ->
+ pt_zipper (NT (prod_lhs prod)) word ->
+ pt_dot
+| Shift_ptd: forall (tok : token) {symbolsq wordq},
+ parse_tree_list symbolsq wordq ->
+ ptl_zipper (T (token_term tok)::symbolsq) (wordq++[tok]) ->
+ pt_dot.
+
+(** We can compute the full semantic value of a parse tree when
+ represented as a dotted ptd. *)
+
+Fixpoint ptlz_sem {hole_symbs hole_word}
+ (ptlz:ptl_zipper hole_symbs hole_word) :
+ (forall A, arrows_right A (map symbol_semantic_type hole_symbs) -> A) ->
+ (symbol_semantic_type (NT (start_nt init))) :=
+ match ptlz with
+ | @Non_terminal_pt_ptlz prod _ ptz =>
+ fun k => ptz_sem ptz (k _ (prod_action prod))
+ | Cons_ptl_ptlz pt ptlz =>
+ fun k => ptlz_sem ptlz (fun _ f => k _ (f (pt_sem pt)))
+ end
+with ptz_sem {hole_symb hole_word}
+ (ptz:pt_zipper hole_symb hole_word):
+ symbol_semantic_type hole_symb -> symbol_semantic_type (NT (start_nt init)) :=
+ match ptz with
+ | Top_ptz => fun sem => sem
+ | Cons_ptl_ptz ptl ptlz =>
+ fun sem => ptlz_sem ptlz (fun _ f => ptl_sem ptl (f sem))
+ end.
+
+Definition ptd_sem (ptd : pt_dot) :=
+ match ptd with
+ | @Reduce_ptd prod _ ptl ptz =>
+ ptz_sem ptz (ptl_sem ptl (prod_action prod))
+ | Shift_ptd tok ptl ptlz =>
+ ptlz_sem ptlz (fun _ f => ptl_sem ptl (f (token_sem tok)))
+ end.
+
+(** The buffer associated with a dotted parse tree corresponds to the
+ buffer left to be read by the parser when at the state represented
+ by the dotted parse tree. *)
+Fixpoint ptlz_buffer {hole_symbs hole_word}
+ (ptlz:ptl_zipper hole_symbs hole_word): buffer :=
+ match ptlz with
+ | Non_terminal_pt_ptlz ptz =>
+ ptz_buffer ptz
+ | @Cons_ptl_ptlz _ _ _ wordt _ ptlz' =>
+ wordt ++ ptlz_buffer ptlz'
+ end
+with ptz_buffer {hole_symb hole_word}
+ (ptz:pt_zipper hole_symb hole_word): buffer :=
+ match ptz with
+ | Top_ptz => buffer_end
+ | Cons_ptl_ptz _ ptlz =>
+ ptlz_buffer ptlz
+ end.
+
+Definition ptd_buffer (ptd:pt_dot) :=
+ match ptd with
+ | Reduce_ptd _ ptz => ptz_buffer ptz
+ | @Shift_ptd tok _ wordq _ ptlz => (tok::ptlz_buffer ptlz)%buf
+ end.
+
+(** We are now ready to define the main invariant of the proof of
+ completeness: we need to specify when a stack is compatible with a
+ dotted parse tree. Informally, a stack is compatible with a dotted
+ parse tree when it is the concatenation stack fragments which are
+ compatible with each of the partially recognized productions
+ appearing in the parse tree zipper. Moreover, the head of each of
+ these stack fragment contains a state which has an item predicted by
+ the corresponding zipper.
+
+ More formally, the compatibility relation first needs the following
+ auxiliary definitions: *)
+Fixpoint ptlz_prod {hole_symbs hole_word}
+ (ptlz:ptl_zipper hole_symbs hole_word): production :=
+ match ptlz with
+ | @Non_terminal_pt_ptlz prod _ _ => prod
+ | Cons_ptl_ptlz _ ptlz' => ptlz_prod ptlz'
+ end.
+
+Fixpoint ptlz_future {hole_symbs hole_word}
+ (ptlz:ptl_zipper hole_symbs hole_word): list symbol :=
+ match ptlz with
+ | Non_terminal_pt_ptlz _ => []
+ | @Cons_ptl_ptlz _ _ s _ _ ptlz' => s::ptlz_future ptlz'
+ end.
+
+Fixpoint ptlz_lookahead {hole_symbs hole_word}
+ (ptlz:ptl_zipper hole_symbs hole_word) : terminal :=
+ match ptlz with
+ | Non_terminal_pt_ptlz ptz => token_term (buf_head (ptz_buffer ptz))
+ | Cons_ptl_ptlz _ ptlz' => ptlz_lookahead ptlz'
+ end.
+
+Fixpoint ptz_stack_compat {hole_symb hole_word}
+ (stk : stack) (ptz : pt_zipper hole_symb hole_word) : Prop :=
+ match ptz with
+ | Top_ptz => stk = []
+ | Cons_ptl_ptz ptl ptlz =>
+ exists stk0,
+ state_has_future (state_of_stack init stk) (ptlz_prod ptlz)
+ (hole_symb::ptlz_future ptlz) (ptlz_lookahead ptlz) /\
+ ptl_stack_compat stk0 ptl stk /\
+ ptlz_stack_compat stk0 ptlz
+ end
+with ptlz_stack_compat {hole_symbs hole_word}
+ (stk : stack) (ptlz : ptl_zipper hole_symbs hole_word) : Prop :=
+ match ptlz with
+ | Non_terminal_pt_ptlz ptz => ptz_stack_compat stk ptz
+ | Cons_ptl_ptlz _ ptlz => ptlz_stack_compat stk ptlz
+ end.
+
+Definition ptd_stack_compat (ptd:pt_dot) (stk:stack): Prop :=
+ match ptd with
+ | @Reduce_ptd prod _ ptl ptz =>
+ exists stk0,
+ state_has_future (state_of_stack init stk) prod []
+ (token_term (buf_head (ptz_buffer ptz))) /\
+ ptl_stack_compat stk0 ptl stk /\
+ ptz_stack_compat stk0 ptz
+ | Shift_ptd tok ptl ptlz =>
+ exists stk0,
+ state_has_future (state_of_stack init stk) (ptlz_prod ptlz)
+ (T (token_term tok) :: ptlz_future ptlz) (ptlz_lookahead ptlz) /\
+ ptl_stack_compat stk0 ptl stk /\
+ ptlz_stack_compat stk0 ptlz
+ end.
+
+Lemma ptz_stack_compat_cons_state_has_future {symbsq wordq symbt wordt} stk
+ (ptl : parse_tree_list symbsq wordq)
+ (ptlz : ptl_zipper (symbt :: symbsq) (wordq ++ wordt)) :
+ ptz_stack_compat stk (Cons_ptl_ptz ptl ptlz) ->
+ state_has_future (state_of_stack init stk) (ptlz_prod ptlz)
+ (symbt::ptlz_future ptlz) (ptlz_lookahead ptlz).
+Proof. move=>[stk0 [? [? ?]]] //. Qed.
+
+Lemma ptlz_future_ptlz_prod hole_symbs hole_word
+ (ptlz:ptl_zipper hole_symbs hole_word) :
+ rev_append (ptlz_future ptlz) hole_symbs = prod_rhs_rev (ptlz_prod ptlz).
+Proof. induction ptlz=>//=. Qed.
+
+Lemma ptlz_future_first {symbs word} (ptlz : ptl_zipper symbs word) :
+ TerminalSet.In (token_term (buf_head (ptlz_buffer ptlz)))
+ (first_word_set (ptlz_future ptlz)) \/
+ token_term (buf_head (ptlz_buffer ptlz)) = ptlz_lookahead ptlz /\
+ nullable_word (ptlz_future ptlz) = true.
+Proof.
+ induction ptlz as [|??? [|tok] pt ptlz IH]; [by auto| |]=>/=.
+ - rewrite (nullable_correct _ _ eq_refl pt).
+ destruct IH as [|[??]]; [left|right]=>/=; auto using TerminalSet.union_3.
+ - left. destruct nullable_symb; eauto using TerminalSet.union_2, first_correct.
+Qed.
+
+(** We now want to define what is the next dotted parse tree which is
+ to be handled after one action. Such dotted parse is built in two
+ steps: Not only we have to perform the action by completing the
+ parse tree, but we also have to prepare for the following step by
+ moving the dot down to place it in front of the next action to be
+ performed.
+*)
+Fixpoint build_pt_dot_from_pt {symb word}
+ (pt : parse_tree symb word) (ptz : pt_zipper symb word)
+ : pt_dot :=
+ match pt in parse_tree symb word
+ return pt_zipper symb word -> pt_dot
+ with
+ | Terminal_pt tok =>
+ fun ptz =>
+ let X :=
+ match ptz in pt_zipper symb word
+ return match symb with T term => True | NT _ => False end ->
+ { symbsq : list symbol &
+ { wordq : list token &
+ (parse_tree_list symbsq wordq *
+ ptl_zipper (symb :: symbsq) (wordq ++ word))%type } }
+ with
+ | Top_ptz => fun F => False_rect _ F
+ | Cons_ptl_ptz ptl ptlz => fun _ =>
+ existT _ _ (existT _ _ (ptl, ptlz))
+ end I
+ in
+ Shift_ptd tok (fst (projT2 (projT2 X))) (snd (projT2 (projT2 X)))
+ | Non_terminal_pt prod ptl => fun ptz =>
+ let is_notnil :=
+ match ptl in parse_tree_list w _
+ return option (match w return Prop with [] => False | _ => True end)
+ with
+ | Nil_ptl => None
+ | _ => Some I
+ end
+ in
+ match is_notnil with
+ | None => Reduce_ptd ptl ptz
+ | Some H => build_pt_dot_from_pt_rec ptl H (Non_terminal_pt_ptlz ptz)
+ end
+ end ptz
+with build_pt_dot_from_pt_rec {symbs word}
+ (ptl : parse_tree_list symbs word)
+ (Hsymbs : match symbs with [] => False | _ => True end)
+ (ptlz : ptl_zipper symbs word)
+ : pt_dot :=
+ match ptl in parse_tree_list symbs word
+ return match symbs with [] => False | _ => True end ->
+ ptl_zipper symbs word ->
+ pt_dot
+ with
+ | Nil_ptl => fun Hsymbs _ => False_rect _ Hsymbs
+ | Cons_ptl ptl' pt => fun _ =>
+ match ptl' in parse_tree_list symbsq wordq
+ return parse_tree_list symbsq wordq ->
+ ptl_zipper (_ :: symbsq) (wordq ++ _) ->
+ pt_dot
+ with
+ | Nil_ptl => fun _ ptlz =>
+ build_pt_dot_from_pt pt (Cons_ptl_ptz Nil_ptl ptlz)
+ | _ => fun ptl' ptlz =>
+ build_pt_dot_from_pt_rec ptl' I (Cons_ptl_ptlz pt ptlz)
+ end ptl'
+ end Hsymbs ptlz.
+
+Definition build_pt_dot_from_ptl {symbs word}
+ (ptl : parse_tree_list symbs word)
+ (ptlz : ptl_zipper symbs word)
+ : pt_dot :=
+ match ptlz in ptl_zipper symbs word
+ return parse_tree_list symbs word -> pt_dot
+ with
+ | Non_terminal_pt_ptlz ptz => fun ptl =>
+ Reduce_ptd ptl ptz
+ | Cons_ptl_ptlz pt ptlz => fun ptl =>
+ build_pt_dot_from_pt pt (Cons_ptl_ptz ptl ptlz)
+ end ptl.
+
+Definition next_ptd (ptd:pt_dot) : option pt_dot :=
+ match ptd with
+ | Shift_ptd tok ptl ptlz =>
+ Some (build_pt_dot_from_ptl (Cons_ptl ptl (Terminal_pt tok)) ptlz)
+ | Reduce_ptd ptl ptz =>
+ match ptz in pt_zipper symb word
+ return parse_tree symb word -> _
+ with
+ | Top_ptz => fun _ => None
+ | Cons_ptl_ptz ptl' ptlz => fun pt =>
+ Some (build_pt_dot_from_ptl (Cons_ptl ptl' pt) ptlz)
+ end (Non_terminal_pt _ ptl)
+ end.
+
+Fixpoint next_ptd_iter (ptd:pt_dot) (log_n_steps:nat) : option pt_dot :=
+ match log_n_steps with
+ | O => next_ptd ptd
+ | S log_n_steps =>
+ match next_ptd_iter ptd log_n_steps with
+ | None => None
+ | Some ptd => next_ptd_iter ptd log_n_steps
+ end
+ end.
+
+(** We prove that these functions behave well w.r.t. semantic values. *)
+Lemma sem_build_from_pt {symb word}
+ (pt : parse_tree symb word) (ptz : pt_zipper symb word) :
+ ptz_sem ptz (pt_sem pt)
+ = ptd_sem (build_pt_dot_from_pt pt ptz)
+with sem_build_from_pt_rec {symbs word}
+ (ptl : parse_tree_list symbs word) (ptlz : ptl_zipper symbs word)
+ Hsymbs :
+ ptlz_sem ptlz (fun _ f => ptl_sem ptl f)
+ = ptd_sem (build_pt_dot_from_pt_rec ptl Hsymbs ptlz).
+Proof.
+ - destruct pt as [tok|prod word ptl]=>/=.
+ + revert ptz. generalize [tok].
+ generalize (token_sem tok). generalize I.
+ change True with (match T (token_term tok) with T _ => True | NT _ => False end) at 1.
+ generalize (T (token_term tok)) => symb HT sem word ptz. by destruct ptz.
+ + match goal with
+ | |- context [match ?X with Some H => _ | None => _ end] => destruct X=>//
+ end.
+ by rewrite -sem_build_from_pt_rec.
+ - destruct ptl; [contradiction|].
+ specialize (sem_build_from_pt_rec _ _ ptl)=>/=. destruct ptl.
+ + by rewrite -sem_build_from_pt.
+ + by rewrite -sem_build_from_pt_rec.
+Qed.
+
+Lemma sem_build_from_ptl {symbs word}
+ (ptl : parse_tree_list symbs word) (ptlz : ptl_zipper symbs word) :
+ ptlz_sem ptlz (fun _ f => ptl_sem ptl f)
+ = ptd_sem (build_pt_dot_from_ptl ptl ptlz).
+Proof. destruct ptlz=>//=. by rewrite -sem_build_from_pt. Qed.
+
+Lemma sem_next_ptd (ptd : pt_dot) :
+ match next_ptd ptd with
+ | None => True
+ | Some ptd' => ptd_sem ptd = ptd_sem ptd'
+ end.
+Proof.
+ destruct ptd as [prod word ptl ptz|tok symbs word ptl ptlz] =>/=.
+ - change (ptl_sem ptl (prod_action prod))
+ with (pt_sem (Non_terminal_pt prod ptl)).
+ generalize (Non_terminal_pt prod ptl). clear ptl.
+ destruct ptz as [|?? ptl ?? ptlz]=>// pt. by rewrite -sem_build_from_ptl.
+ - by rewrite -sem_build_from_ptl.
+Qed.
+
+Lemma sem_next_ptd_iter (ptd : pt_dot) (log_n_steps : nat) :
+ match next_ptd_iter ptd log_n_steps with
+ | None => True
+ | Some ptd' => ptd_sem ptd = ptd_sem ptd'
+ end.
+Proof.
+ revert ptd.
+ induction log_n_steps as [|log_n_steps IH]; [by apply sem_next_ptd|]=>/= ptd.
+ assert (IH1 := IH ptd). destruct next_ptd_iter as [ptd'|]=>//.
+ specialize (IH ptd'). destruct next_ptd_iter=>//. congruence.
+Qed.
+
+(** We prove that these functions behave well w.r.t. xxx_buffer. *)
+Lemma ptd_buffer_build_from_pt {symb word}
+ (pt : parse_tree symb word) (ptz : pt_zipper symb word) :
+ (word ++ ptz_buffer ptz)%buf = ptd_buffer (build_pt_dot_from_pt pt ptz)
+with ptd_buffer_build_from_pt_rec {symbs word}
+ (ptl : parse_tree_list symbs word) (ptlz : ptl_zipper symbs word)
+ Hsymbs :
+ (word ++ ptlz_buffer ptlz)%buf = ptd_buffer (build_pt_dot_from_pt_rec ptl Hsymbs ptlz).
+Proof.
+ - destruct pt as [tok|prod word ptl]=>/=.
+ + f_equal. revert ptz. generalize [tok].
+ generalize (token_sem tok). generalize I.
+ change True with (match T (token_term tok) with T _ => True | NT _ => False end) at 1.
+ generalize (T (token_term tok)) => symb HT sem word ptz. by destruct ptz.
+ + match goal with
+ | |- context [match ?X with Some H => _ | None => _ end] => destruct X eqn:EQ
+ end.
+ * by rewrite -ptd_buffer_build_from_pt_rec.
+ * rewrite [X in (X ++ _)%buf](_ : word = []) //. clear -EQ. by destruct ptl.
+ - destruct ptl as [|?? ptl ?? pt]; [contradiction|].
+ specialize (ptd_buffer_build_from_pt_rec _ _ ptl).
+ destruct ptl.
+ + by rewrite /= -ptd_buffer_build_from_pt.
+ + by rewrite -ptd_buffer_build_from_pt_rec //= app_buf_assoc.
+Qed.
+
+Lemma ptd_buffer_build_from_ptl {symbs word}
+ (ptl : parse_tree_list symbs word) (ptlz : ptl_zipper symbs word) :
+ ptlz_buffer ptlz = ptd_buffer (build_pt_dot_from_ptl ptl ptlz).
+Proof.
+ destruct ptlz as [|???? pt]=>//=. by rewrite -ptd_buffer_build_from_pt.
+Qed.
+
+(** We prove that these functions behave well w.r.t. xxx_stack_compat. *)
+Lemma ptd_stack_compat_build_from_pt {symb word}
+ (pt : parse_tree symb word) (ptz : pt_zipper symb word)
+ (stk: stack) :
+ ptz_stack_compat stk ptz ->
+ ptd_stack_compat (build_pt_dot_from_pt pt ptz) stk
+with ptd_stack_compat_build_from_pt_rec {symbs word}
+ (ptl : parse_tree_list symbs word) (ptlz : ptl_zipper symbs word)
+ (stk : stack) Hsymbs :
+ ptlz_stack_compat stk ptlz ->
+ state_has_future (state_of_stack init stk) (ptlz_prod ptlz)
+ (rev' (prod_rhs_rev (ptlz_prod ptlz))) (ptlz_lookahead ptlz) ->
+ ptd_stack_compat (build_pt_dot_from_pt_rec ptl Hsymbs ptlz) stk.
+Proof.
+ - intros Hstk. destruct pt as [tok|prod word ptl]=>/=.
+ + revert ptz Hstk. generalize [tok]. generalize (token_sem tok). generalize I.
+ change True with (match T (token_term tok) with T _ => True | NT _ => False end) at 1.
+ generalize (T (token_term tok)) => symb HT sem word ptz. by destruct ptz.
+ + assert (state_has_future (state_of_stack init stk) prod
+ (rev' (prod_rhs_rev prod)) (token_term (buf_head (ptz_buffer ptz)))).
+ { revert ptz Hstk. remember (NT (prod_lhs prod)) eqn:EQ=>ptz.
+ destruct ptz as [|?? ptl0 ?? ptlz0].
+ - intros ->. apply start_future. congruence.
+ - subst. intros (stk0 & Hfut & _). apply non_terminal_closed in Hfut.
+ specialize (Hfut prod eq_refl).
+ destruct (ptlz_future_first ptlz0) as [Hfirst|[Hfirst Hnull]].
+ + destruct Hfut as [_ Hfut]. auto.
+ + destruct Hfut as [Hfut _]. by rewrite Hnull -Hfirst in Hfut. }
+ match goal with
+ | |- context [match ?X with Some H => _ | None => _ end] => destruct X eqn:EQ
+ end.
+ * by apply ptd_stack_compat_build_from_pt_rec.
+ * exists stk. destruct ptl=>//.
+ - intros Hstk Hfut. destruct ptl as [|?? ptl ?? pt]; [contradiction|].
+ specialize (ptd_stack_compat_build_from_pt_rec _ _ ptl). destruct ptl.
+ + eapply ptd_stack_compat_build_from_pt=>//. exists stk.
+ split; [|split]=>//; [].
+ by rewrite -ptlz_future_ptlz_prod rev_append_rev /rev' -rev_alt
+ rev_app_distr rev_involutive in Hfut.
+ + by apply ptd_stack_compat_build_from_pt_rec.
+Qed.
+
+Lemma ptd_stack_compat_build_from_ptl {symbs word}
+ (ptl : parse_tree_list symbs word) (ptlz : ptl_zipper symbs word)
+ (stk stk0: stack) :
+ ptlz_stack_compat stk0 ptlz ->
+ ptl_stack_compat stk0 ptl stk ->
+ state_has_future (state_of_stack init stk) (ptlz_prod ptlz)
+ (ptlz_future ptlz) (ptlz_lookahead ptlz) ->
+ ptd_stack_compat (build_pt_dot_from_ptl ptl ptlz) stk.
+Proof.
+ intros Hstk0 Hstk Hfut. destruct ptlz=>/=.
+ - eauto.
+ - apply ptd_stack_compat_build_from_pt=>/=. eauto.
+Qed.
+
+(** We can now proceed by proving that the invariant is preserved by
+ each step of parsing. We also prove that each step of parsing
+ follows next_ptd.
+
+ We start with reduce steps: *)
+Lemma reduce_step_next_ptd (prod : production) (word : list token)
+ (ptl : parse_tree_list (prod_rhs_rev prod) word)
+ (ptz : pt_zipper (NT (prod_lhs prod)) word)
+ (stk : stack)
+ Hval Hi :
+ ptd_stack_compat (Reduce_ptd ptl ptz) stk ->
+ match next_ptd (Reduce_ptd ptl ptz) with
+ | None =>
+ reduce_step init stk prod (ptz_buffer ptz) Hval Hi =
+ Accept_sr (ptd_sem (Reduce_ptd ptl ptz)) buffer_end
+ | Some ptd =>
+ exists stk',
+ reduce_step init stk prod (ptz_buffer ptz) Hval Hi =
+ Progress_sr stk' (ptd_buffer ptd) /\
+ ptd_stack_compat ptd stk'
+ end.
+Proof.
+ intros (stk0 & _ & Hstk & Hstk0).
+ apply pop_stack_compat_pop_spec with (action := prod_action prod) in Hstk.
+ rewrite <-pop_spec_ok with (Hp := reduce_step_subproof init stk prod Hval Hi) in Hstk.
+ unfold reduce_step.
+ match goal with
+ | |- context [pop_state_valid init ?A stk ?B ?C ?D ?E ?F] =>
+ generalize (pop_state_valid init A stk B C D E F)
+ end.
+ rewrite Hstk /=. intros Hv.
+ generalize (reduce_step_subproof1 init stk prod Hval stk0 (fun _ : True => Hv)).
+ clear Hval Hstk Hi Hv stk.
+ assert (Hgoto := fun fut prod' =>
+ non_terminal_goto (state_of_stack init stk0) prod' (NT (prod_lhs prod)::fut)).
+ simpl in Hgoto.
+ destruct goto_table as [[st Hst]|] eqn:Hgoto'.
+ - intros _.
+ assert (match ptz with Top_ptz => False | _ => True end).
+ { revert ptz Hst Hstk0 Hgoto'.
+ generalize (eq_refl (NT (prod_lhs prod))).
+ generalize (NT (prod_lhs prod)) at 1 3 5.
+ intros nt Hnt ptz. destruct ptz=>//. injection Hnt=> <- /= Hst -> /= Hg.
+ assert (Hsg := start_goto init). by rewrite Hg in Hsg. }
+ clear Hgoto'.
+
+ change (ptl_sem ptl (prod_action prod))
+ with (pt_sem (Non_terminal_pt prod ptl)).
+ generalize (Non_terminal_pt prod ptl). clear ptl.
+ destruct ptz as [|?? ptl ? ? ptlz]=>// pt.
+
+ subst=>/=. eexists _. split.
+ + f_equal. apply ptd_buffer_build_from_ptl.
+ + destruct Hstk0 as (stk0' & Hfut & Hstk0' & Hstk0).
+ apply (ptd_stack_compat_build_from_ptl _ _ _ stk0'); auto; [].
+ split=>//. by exists eq_refl.
+ - intros Hv. generalize (reduce_step_subproof0 _ prod _ (fun _ => Hv)).
+ intros EQnt. clear Hv Hgoto'.
+
+ change (ptl_sem ptl (prod_action prod))
+ with (pt_sem (Non_terminal_pt prod ptl)).
+ generalize (Non_terminal_pt prod ptl). clear ptl. destruct ptz.
+ + intros pt. f_equal. by rewrite cast_eq.
+ + edestruct Hgoto. eapply ptz_stack_compat_cons_state_has_future, Hstk0.
+Qed.
+
+Lemma step_next_ptd (ptd : pt_dot) (stk : stack) Hi :
+ ptd_stack_compat ptd stk ->
+ match next_ptd ptd with
+ | None =>
+ step safe init stk (ptd_buffer ptd) Hi =
+ Accept_sr (ptd_sem ptd) buffer_end
+ | Some ptd' =>
+ exists stk',
+ step safe init stk (ptd_buffer ptd) Hi =
+ Progress_sr stk' (ptd_buffer ptd') /\
+ ptd_stack_compat ptd' stk'
+ end.
+Proof.
+ intros Hstk. unfold step.
+ generalize (reduce_ok safe (state_of_stack init stk)).
+ destruct ptd as [prod word ptl ptz|tok symbs word ptl ptlz].
+ - assert (Hfut : state_has_future (state_of_stack init stk) prod []
+ (token_term (buf_head (ptz_buffer ptz)))).
+ { destruct Hstk as (? & ? & ?)=>//. }
+ assert (Hact := end_reduce _ _ _ _ Hfut).
+ destruct action_table as [?|awt]=>Hval /=.
+ + subst. by apply reduce_step_next_ptd.
+ + set (term := token_term (buf_head (ptz_buffer ptz))) in *.
+ generalize (Hval term). clear Hval. destruct (awt term)=>//. subst.
+ intros Hval. by apply reduce_step_next_ptd.
+ - destruct Hstk as (stk0 & Hfut & Hstk & Hstk0).
+ assert (Hact := terminal_shift _ _ _ _ Hfut). simpl in Hact. clear Hfut.
+ destruct action_table as [?|awt]=>//= /(_ (token_term tok)).
+ destruct awt as [st' EQ| |]=>// _. eexists. split.
+ + f_equal. rewrite -ptd_buffer_build_from_ptl //.
+ + apply (ptd_stack_compat_build_from_ptl _ _ _ stk0); simpl; eauto.
+Qed.
+
+(** We prove the completeness of the parser main loop. *)
+Lemma parse_fix_next_ptd_iter (ptd : pt_dot) (stk : stack) (log_n_steps : nat) Hi :
+ ptd_stack_compat ptd stk ->
+ match next_ptd_iter ptd log_n_steps with
+ | None =>
+ proj1_sig (parse_fix safe init stk (ptd_buffer ptd) log_n_steps Hi) =
+ Accept_sr (ptd_sem ptd) buffer_end
+ | Some ptd' =>
+ exists stk',
+ proj1_sig (parse_fix safe init stk (ptd_buffer ptd) log_n_steps Hi) =
+ Progress_sr stk' (ptd_buffer ptd') /\
+ ptd_stack_compat ptd' stk'
+ end.
+Proof.
+ revert ptd stk Hi.
+ induction log_n_steps as [|log_n_steps IH]; [by apply step_next_ptd|].
+ move => /= ptd stk Hi Hstk. assert (IH1 := IH ptd stk Hi Hstk).
+ assert (EQsem := sem_next_ptd_iter ptd log_n_steps).
+ destruct parse_fix as [sr Hi']. simpl in IH1.
+ destruct next_ptd_iter as [ptd'|].
+ - rewrite EQsem. destruct IH1 as (stk' & -> & Hstk'). by apply IH.
+ - by subst.
+Qed.
+
+(** The parser is defined by recursion over a fuel parameter. In the
+ completeness proof, we need to predict how much fuel is going to be
+ needed in order to prove that enough fuel gives rise to a successful
+ parsing.
+
+ To do so, of a dotted parse tree, which is the number of actions
+ left to be executed before complete parsing when the current state
+ is represented by the dotted parse tree. *)
+Fixpoint ptlz_cost {hole_symbs hole_word}
+ (ptlz:ptl_zipper hole_symbs hole_word) :=
+ match ptlz with
+ | Non_terminal_pt_ptlz ptz => ptz_cost ptz
+ | Cons_ptl_ptlz pt ptlz' => pt_size pt + ptlz_cost ptlz'
+ end
+with ptz_cost {hole_symb hole_word} (ptz:pt_zipper hole_symb hole_word) :=
+ match ptz with
+ | Top_ptz => 0
+ | Cons_ptl_ptz ptl ptlz' => 1 + ptlz_cost ptlz'
+ end.
+
+Definition ptd_cost (ptd:pt_dot) :=
+ match ptd with
+ | Reduce_ptd ptl ptz => ptz_cost ptz
+ | Shift_ptd _ ptl ptlz => 1 + ptlz_cost ptlz
+ end.
+
+Lemma ptd_cost_build_from_pt {symb word}
+ (pt : parse_tree symb word) (ptz : pt_zipper symb word) :
+ pt_size pt + ptz_cost ptz = S (ptd_cost (build_pt_dot_from_pt pt ptz))
+with ptd_cost_build_from_pt_rec {symbs word}
+ (ptl : parse_tree_list symbs word) (ptlz : ptl_zipper symbs word)
+ Hsymbs :
+ ptl_size ptl + ptlz_cost ptlz = ptd_cost (build_pt_dot_from_pt_rec ptl Hsymbs ptlz).
+Proof.
+ - destruct pt as [tok|prod word ptl']=>/=.
+ + revert ptz. generalize [tok]. generalize (token_sem tok). generalize I.
+ change True with (match T (token_term tok) with T _ => True | NT _ => False end) at 1.
+ generalize (T (token_term tok)) => symb HT sem word ptz. by destruct ptz.
+ + match goal with
+ | |- context [match ?X with Some H => _ | None => _ end] => destruct X eqn:EQ
+ end.
+ * rewrite -ptd_cost_build_from_pt_rec /= plus_n_Sm //.
+ * simpl. by destruct ptl'.
+ - destruct ptl as [|?? ptl ?? pt]; [contradiction|].
+ specialize (ptd_cost_build_from_pt_rec _ _ ptl). destruct ptl.
+ + apply eq_add_S. rewrite -ptd_cost_build_from_pt /=. ring.
+ + rewrite -ptd_cost_build_from_pt_rec //=. ring.
+Qed.
+
+Lemma ptd_cost_build_from_ptl {symbs word}
+ (ptl : parse_tree_list symbs word) (ptlz : ptl_zipper symbs word) :
+ ptlz_cost ptlz = ptd_cost (build_pt_dot_from_ptl ptl ptlz).
+Proof.
+ destruct ptlz=>//. apply eq_add_S. rewrite -ptd_cost_build_from_pt /=. ring.
+Qed.
+
+Lemma next_ptd_cost ptd:
+ match next_ptd ptd with
+ | None => ptd_cost ptd = 0
+ | Some ptd' => ptd_cost ptd = S (ptd_cost ptd')
+ end.
+Proof.
+ destruct ptd as [prod word ptl ptz|tok symbq wordq ptl ptlz] =>/=.
+ - generalize (Non_terminal_pt prod ptl). clear ptl.
+ destruct ptz as [|?? ptl ?? ptlz]=>// pt. by rewrite -ptd_cost_build_from_ptl.
+ - by rewrite -ptd_cost_build_from_ptl.
+Qed.
+
+Lemma next_ptd_iter_cost ptd log_n_steps :
+ match next_ptd_iter ptd log_n_steps with
+ | None => ptd_cost ptd < 2^log_n_steps
+ | Some ptd' => ptd_cost ptd = 2^log_n_steps + ptd_cost ptd'
+ end.
+Proof.
+ revert ptd. induction log_n_steps as [|log_n_steps IH]=>ptd /=.
+ - assert (Hptd := next_ptd_cost ptd). destruct next_ptd=>//. by rewrite Hptd.
+ - rewrite Nat.add_0_r. assert (IH1 := IH ptd). destruct next_ptd_iter as [ptd'|].
+ + specialize (IH ptd'). destruct next_ptd_iter as [ptd''|].
+ * by rewrite IH1 IH -!plus_assoc.
+ * rewrite IH1. by apply plus_lt_compat_l.
+ + by apply lt_plus_trans.
+Qed.
+
+(** We now prove the top-level parsing function. The only thing that
+ is left to be done is the initialization. To do so, we define the
+ initial dotted parse tree, depending on a full (top-level) parse tree. *)
+
+Variable full_pt : parse_tree (NT (start_nt init)) full_word.
+
+Theorem parse_complete log_n_steps:
+ match parse safe init (full_word ++ buffer_end) log_n_steps with
+ | Parsed_pr sem buff =>
+ sem = pt_sem full_pt /\ buff = buffer_end /\ pt_size full_pt <= 2^log_n_steps
+ | Timeout_pr => 2^log_n_steps < pt_size full_pt
+ | Fail_pr => False
+ end.
+Proof.
+ assert (Hstk : ptd_stack_compat (build_pt_dot_from_pt full_pt Top_ptz) []) by
+ by apply ptd_stack_compat_build_from_pt.
+ unfold parse.
+ assert (Hparse := parse_fix_next_ptd_iter _ _ log_n_steps (parse_subproof init) Hstk).
+ rewrite -ptd_buffer_build_from_pt -sem_build_from_pt /= in Hparse.
+ assert (Hcost := next_ptd_iter_cost (build_pt_dot_from_pt full_pt Top_ptz) log_n_steps).
+ destruct next_ptd_iter.
+ - destruct Hparse as (? & -> & ?). apply (f_equal S) in Hcost.
+ rewrite -ptd_cost_build_from_pt Nat.add_0_r in Hcost. rewrite Hcost.
+ apply le_lt_n_Sm, le_plus_l.
+ - rewrite Hparse. split; [|split]=>//. apply lt_le_S in Hcost.
+ by rewrite -ptd_cost_build_from_pt Nat.add_0_r in Hcost.
+Qed.
+
+End Completeness_Proof.
+
+End Make.
diff --git a/MenhirLib/Interpreter_correct.v b/MenhirLib/Interpreter_correct.v
new file mode 100644
index 00000000..1325f610
--- /dev/null
+++ b/MenhirLib/Interpreter_correct.v
@@ -0,0 +1,175 @@
+(****************************************************************************)
+(* *)
+(* Menhir *)
+(* *)
+(* Jacques-Henri Jourdan, CNRS, LRI, Université Paris Sud *)
+(* *)
+(* Copyright Inria. All rights reserved. This file is distributed under *)
+(* the terms of the GNU Lesser General Public License as published by the *)
+(* Free Software Foundation, either version 3 of the License, or (at your *)
+(* option) any later version, as described in the file LICENSE. *)
+(* *)
+(****************************************************************************)
+
+From Coq Require Import List Syntax.
+Require Import Alphabet.
+Require Grammar Automaton Interpreter.
+From Coq.ssr Require Import ssreflect.
+
+Module Make(Import A:Automaton.T) (Import Inter:Interpreter.T A).
+
+(** * Correctness of the interpreter **)
+
+(** We prove that, in any case, if the interpreter accepts returning a
+ semantic value, then this is a semantic value of the input **)
+
+Section Init.
+
+Variable init:initstate.
+
+(** [word_has_stack_semantics] relates a word with a stack, stating that the
+ word is a concatenation of words that have the semantic values stored in
+ the stack. **)
+Inductive word_has_stack_semantics:
+ forall (word:list token) (stack:stack), Prop :=
+ | Nil_stack_whss: word_has_stack_semantics [] []
+ | Cons_stack_whss:
+ forall (wordq:list token) (stackq:stack),
+ word_has_stack_semantics wordq stackq ->
+
+ forall (wordt:list token) (s:noninitstate)
+ (pt:parse_tree (last_symb_of_non_init_state s) wordt),
+
+ word_has_stack_semantics
+ (wordq++wordt) (existT noninitstate_type s (pt_sem pt)::stackq).
+
+(** [pop] preserves the invariant **)
+Lemma pop_spec_ptl A symbols_to_pop action word_stk stk (res : A) stk' :
+ pop_spec symbols_to_pop stk action stk' res ->
+ word_has_stack_semantics word_stk stk ->
+ exists word_stk' word_res (ptl:parse_tree_list symbols_to_pop word_res),
+ (word_stk' ++ word_res = word_stk)%list /\
+ word_has_stack_semantics word_stk' stk' /\
+ ptl_sem ptl action = res.
+Proof.
+ intros Hspec. revert word_stk.
+ induction Hspec as [stk sem|symbols_to_pop st stk action sem stk' res Hspec IH];
+ intros word_stk Hword_stk.
+ - exists word_stk, [], Nil_ptl. rewrite -app_nil_end. eauto.
+ - inversion Hword_stk. subst_existT.
+ edestruct IH as (word_stk' & word_res & ptl & ? & Hword_stk'' & ?); [eassumption|].
+ subst. eexists word_stk', (word_res ++ _)%list, (Cons_ptl ptl _).
+ split; [|split]=>//. rewrite app_assoc //.
+Qed.
+
+(** [reduce_step] preserves the invariant **)
+Lemma reduce_step_invariant (stk:stack) (prod:production) Hv Hi word buffer :
+ word_has_stack_semantics word stk ->
+ match reduce_step init stk prod buffer Hv Hi with
+ | Accept_sr sem buffer_new =>
+ exists pt : parse_tree (NT (start_nt init)) word,
+ buffer = buffer_new /\ pt_sem pt = sem
+ | Progress_sr stk' buffer_new =>
+ buffer = buffer_new /\ word_has_stack_semantics word stk'
+ | Fail_sr => True
+ end.
+Proof.
+ intros Hword_stk. unfold reduce_step.
+ match goal with
+ | |- context [pop_state_valid init ?stp stk ?x1 ?x2 ?x3 ?x4 ?x5] =>
+ generalize (pop_state_valid init stp stk x1 x2 x3 x4 x5)
+ end.
+ destruct pop as [stk' sem] eqn:Hpop=>/= Hv'.
+ apply pop_spec_ok in Hpop. apply pop_spec_ptl with (word_stk := word) in Hpop=>//.
+ destruct Hpop as (word1 & word2 & ptl & <- & Hword1 & <-).
+ generalize (reduce_step_subproof1 init stk prod Hv stk' (fun _ : True => Hv')).
+ destruct goto_table as [[st' EQ]|].
+ - intros _. split=>//.
+ change (ptl_sem ptl (prod_action prod)) with (pt_sem (Non_terminal_pt prod ptl)).
+ generalize (Non_terminal_pt prod ptl). rewrite ->EQ. intros pt. by constructor.
+ - intros Hstk'. destruct Hword1; [|by destruct Hstk'].
+ generalize (reduce_step_subproof0 init prod [] (fun _ : True => Hstk')).
+ simpl in Hstk'. rewrite -Hstk' // => EQ. rewrite cast_eq.
+ exists (Non_terminal_pt prod ptl). by split.
+Qed.
+
+(** [step] preserves the invariant **)
+Lemma step_invariant stk word buffer safe Hi :
+ word_has_stack_semantics word stk ->
+ match step safe init stk buffer Hi with
+ | Accept_sr sem buffer_new =>
+ exists word_new (pt:parse_tree (NT (start_nt init)) word_new),
+ (word ++ buffer = word_new ++ buffer_new)%buf /\
+ pt_sem pt = sem
+ | Progress_sr stk_new buffer_new =>
+ exists word_new,
+ (word ++ buffer = word_new ++ buffer_new)%buf /\
+ word_has_stack_semantics word_new stk_new
+ | Fail_sr => True
+ end.
+Proof.
+ intros Hword_stk. unfold step.
+ generalize (reduce_ok safe (state_of_stack init stk)).
+ destruct action_table as [prod|awt].
+ - intros Hv.
+ apply (reduce_step_invariant stk prod (fun _ => Hv) Hi word buffer) in Hword_stk.
+ destruct reduce_step=>//.
+ + destruct Hword_stk as (pt & <- & <-); eauto.
+ + destruct Hword_stk as [<- ?]; eauto.
+ - destruct buffer as [tok buffer]=>/=.
+ move=> /(_ (token_term tok)) Hv. destruct (awt (token_term tok)) as [st EQ|prod|]=>//.
+ + eexists _. split; [by apply app_buf_assoc with (l2 := [_])|].
+ change (token_sem tok) with (pt_sem (Terminal_pt tok)).
+ generalize (Terminal_pt tok). generalize [tok].
+ rewrite -> EQ=>word' pt /=. by constructor.
+ + apply (reduce_step_invariant stk prod (fun _ => Hv) Hi word (tok::buffer))
+ in Hword_stk.
+ destruct reduce_step=>//.
+ * destruct Hword_stk as (pt & <- & <-); eauto.
+ * destruct Hword_stk as [<- ?]; eauto.
+Qed.
+
+(** [step] preserves the invariant **)
+Lemma parse_fix_invariant stk word buffer safe log_n_steps Hi :
+ word_has_stack_semantics word stk ->
+ match proj1_sig (parse_fix safe init stk buffer log_n_steps Hi) with
+ | Accept_sr sem buffer_new =>
+ exists word_new (pt:parse_tree (NT (start_nt init)) word_new),
+ (word ++ buffer = word_new ++ buffer_new)%buf /\
+ pt_sem pt = sem
+ | Progress_sr stk_new buffer_new =>
+ exists word_new,
+ (word ++ buffer = word_new ++ buffer_new)%buf /\
+ word_has_stack_semantics word_new stk_new
+ | Fail_sr => True
+ end.
+Proof.
+ revert stk word buffer Hi.
+ induction log_n_steps as [|log_n_steps IH]=>/= stk word buffer Hi Hstk;
+ [by apply step_invariant|].
+ assert (IH1 := IH stk word buffer Hi Hstk).
+ destruct parse_fix as [[] Hi']=>/=; try by apply IH1.
+ destruct IH1 as (word' & -> & Hstk')=>//. by apply IH.
+Qed.
+
+(** The interpreter is correct : if it returns a semantic value, then the input
+ word has this semantic value.
+**)
+Theorem parse_correct safe buffer log_n_steps:
+ match parse safe init buffer log_n_steps with
+ | Parsed_pr sem buffer_new =>
+ exists word_new (pt:parse_tree (NT (start_nt init)) word_new),
+ buffer = (word_new ++ buffer_new)%buf /\
+ pt_sem pt = sem
+ | _ => True
+ end.
+Proof.
+ unfold parse.
+ assert (Hparse := parse_fix_invariant [] [] buffer safe log_n_steps
+ (parse_subproof init)).
+ destruct proj1_sig=>//. apply Hparse. constructor.
+Qed.
+
+End Init.
+
+End Make.
diff --git a/MenhirLib/Main.v b/MenhirLib/Main.v
new file mode 100644
index 00000000..f6158074
--- /dev/null
+++ b/MenhirLib/Main.v
@@ -0,0 +1,79 @@
+(****************************************************************************)
+(* *)
+(* Menhir *)
+(* *)
+(* Jacques-Henri Jourdan, CNRS, LRI, Université Paris Sud *)
+(* *)
+(* Copyright Inria. All rights reserved. This file is distributed under *)
+(* the terms of the GNU Lesser General Public License as published by the *)
+(* Free Software Foundation, either version 3 of the License, or (at your *)
+(* option) any later version, as described in the file LICENSE. *)
+(* *)
+(****************************************************************************)
+
+Require Grammar Automaton Interpreter_correct Interpreter_complete.
+From Coq Require Import Syntax Arith.
+
+Module Make(Export Aut:Automaton.T).
+Export Aut.Gram.
+Export Aut.GramDefs.
+
+Module Import Inter := Interpreter.Make Aut.
+Module Correct := Interpreter_correct.Make Aut Inter.
+Module Complete := Interpreter_complete.Make Aut Inter.
+
+Definition complete_validator:unit->bool := Complete.Valid.is_complete.
+Definition safe_validator:unit->bool := ValidSafe.is_safe.
+Definition parse (safe:safe_validator ()=true) init log_n_steps buffer :
+ parse_result (symbol_semantic_type (NT (start_nt init))):=
+ parse (ValidSafe.safe_is_validator safe) init buffer log_n_steps.
+
+(** Correction theorem. **)
+Theorem parse_correct
+ (safe:safe_validator ()= true) init log_n_steps buffer:
+ match parse safe init log_n_steps buffer with
+ | Parsed_pr sem buffer_new =>
+ exists word (pt : parse_tree (NT (start_nt init)) word),
+ buffer = (word ++ buffer_new)%buf /\
+ pt_sem pt = sem
+ | _ => True
+ end.
+Proof. apply Correct.parse_correct. Qed.
+
+(** Completeness theorem. **)
+Theorem parse_complete
+ (safe:safe_validator () = true) init log_n_steps word buffer_end:
+ complete_validator () = true ->
+ forall tree:parse_tree (NT (start_nt init)) word,
+ match parse safe init log_n_steps (word ++ buffer_end) with
+ | Fail_pr => False
+ | Parsed_pr sem_res buffer_end_res =>
+ sem_res = pt_sem tree /\ buffer_end_res = buffer_end /\
+ pt_size tree <= 2^log_n_steps
+ | Timeout_pr => 2^log_n_steps < pt_size tree
+ end.
+Proof.
+ intros. now apply Complete.parse_complete, Complete.Valid.complete_is_validator.
+Qed.
+
+(** Unambiguity theorem. **)
+Theorem unambiguity:
+ safe_validator () = true -> complete_validator () = true -> inhabited token ->
+ forall init word,
+ forall (tree1 tree2:parse_tree (NT (start_nt init)) word),
+ pt_sem tree1 = pt_sem tree2.
+Proof.
+ intros Hsafe Hcomp [tok] init word tree1 tree2.
+ pose (buf_end := cofix buf_end := (tok :: buf_end)%buf).
+ assert (Hcomp1 := parse_complete Hsafe init (pt_size tree1) word buf_end
+ Hcomp tree1).
+ assert (Hcomp2 := parse_complete Hsafe init (pt_size tree1) word buf_end
+ Hcomp tree2).
+ destruct parse.
+ - destruct Hcomp1.
+ - exfalso. eapply PeanoNat.Nat.lt_irrefl. etransitivity; [|apply Hcomp1].
+ eapply Nat.pow_gt_lin_r. constructor.
+ - destruct Hcomp1 as [-> _], Hcomp2 as [-> _]. reflexivity.
+Qed.
+
+End Make.
diff --git a/MenhirLib/Validator_classes.v b/MenhirLib/Validator_classes.v
new file mode 100644
index 00000000..d8063123
--- /dev/null
+++ b/MenhirLib/Validator_classes.v
@@ -0,0 +1,75 @@
+(****************************************************************************)
+(* *)
+(* Menhir *)
+(* *)
+(* Jacques-Henri Jourdan, CNRS, LRI, Université Paris Sud *)
+(* *)
+(* Copyright Inria. All rights reserved. This file is distributed under *)
+(* the terms of the GNU Lesser General Public License as published by the *)
+(* Free Software Foundation, either version 3 of the License, or (at your *)
+(* option) any later version, as described in the file LICENSE. *)
+(* *)
+(****************************************************************************)
+
+From Coq Require Import List.
+From Coq.ssr Require Import ssreflect.
+Require Import Alphabet.
+
+Class IsValidator (P : Prop) (b : bool) :=
+ is_validator : b = true -> P.
+Hint Mode IsValidator + - : typeclass_instances.
+
+Instance is_validator_true : IsValidator True true.
+Proof. done. Qed.
+
+Instance is_validator_false : IsValidator False false.
+Proof. done. Qed.
+
+Instance is_validator_eq_true b :
+ IsValidator (b = true) b.
+Proof. done. Qed.
+
+Instance is_validator_and P1 b1 P2 b2 `{IsValidator P1 b1} `{IsValidator P2 b2}:
+ IsValidator (P1 /\ P2) (if b1 then b2 else false).
+Proof. by split; destruct b1, b2; apply is_validator. Qed.
+
+Instance is_validator_comparable_leibniz_eq A (C:Comparable A) (x y : A) :
+ ComparableLeibnizEq C ->
+ IsValidator (x = y) (compare_eqb x y).
+Proof. intros ??. by apply compare_eqb_iff. Qed.
+
+Instance is_validator_comparable_eq_impl A `(Comparable A) (x y : A) P b :
+ IsValidator P b ->
+ IsValidator (x = y -> P) (if compare_eqb x y then b else true).
+Proof.
+ intros Hval Val ->. rewrite /compare_eqb compare_refl in Val. auto.
+Qed.
+
+Lemma is_validator_forall_finite A P b `(Finite A) :
+ (forall (x : A), IsValidator (P x) (b x)) ->
+ IsValidator (forall (x : A), P x) (forallb b all_list).
+Proof.
+ move=> ? /forallb_forall Hb ?.
+ apply is_validator, Hb, all_list_forall.
+Qed.
+
+(* We do not use an instance directly here, because we need somehow to
+ force Coq to instantiate b with a lambda. *)
+Hint Extern 2 (IsValidator (forall x : ?A, _) _) =>
+ eapply (is_validator_forall_finite _ _ (fun (x:A) => _))
+ : typeclass_instances.
+
+(* Hint for synthetizing pattern-matching. *)
+Hint Extern 2 (IsValidator (match ?u with _ => _ end) ?b0) =>
+ let b := fresh "b" in
+ unshelve notypeclasses refine (let b : bool := _ in _);
+ [destruct u; intros; shelve|]; (* Synthetize `match .. with` in the validator. *)
+ unify b b0;
+ unfold b; destruct u; clear b
+ : typeclass_instances.
+
+(* Hint for unfolding definitions. This is necessary because many
+ hints for IsValidator use [Hint Extern], which do not automatically
+ unfold identifiers. *)
+Hint Extern 100 (IsValidator ?X _) => unfold X
+ : typeclass_instances.
diff --git a/MenhirLib/Validator_complete.v b/MenhirLib/Validator_complete.v
new file mode 100644
index 00000000..ebb74500
--- /dev/null
+++ b/MenhirLib/Validator_complete.v
@@ -0,0 +1,394 @@
+(****************************************************************************)
+(* *)
+(* Menhir *)
+(* *)
+(* Jacques-Henri Jourdan, CNRS, LRI, Université Paris Sud *)
+(* *)
+(* Copyright Inria. All rights reserved. This file is distributed under *)
+(* the terms of the GNU Lesser General Public License as published by the *)
+(* Free Software Foundation, either version 3 of the License, or (at your *)
+(* option) any later version, as described in the file LICENSE. *)
+(* *)
+(****************************************************************************)
+
+From Coq Require Import List Syntax Derive.
+From Coq.ssr Require Import ssreflect.
+Require Automaton.
+Require Import Alphabet Validator_classes.
+
+Module Make(Import A:Automaton.T).
+
+(** We instantiate some sets/map. **)
+Module TerminalComparableM <: ComparableM.
+ Definition t := terminal.
+ Instance tComparable : Comparable t := _.
+End TerminalComparableM.
+Module TerminalOrderedType := OrderedType_from_ComparableM TerminalComparableM.
+Module StateProdPosComparableM <: ComparableM.
+ Definition t := (state*production*nat)%type.
+ Instance tComparable : Comparable t := _.
+End StateProdPosComparableM.
+Module StateProdPosOrderedType :=
+ OrderedType_from_ComparableM StateProdPosComparableM.
+
+Module TerminalSet := FSetAVL.Make TerminalOrderedType.
+Module StateProdPosMap := FMapAVL.Make StateProdPosOrderedType.
+
+(** Nullable predicate for symbols and list of symbols. **)
+Definition nullable_symb (symbol:symbol) :=
+ match symbol with
+ | NT nt => nullable_nterm nt
+ | _ => false
+ end.
+
+Definition nullable_word (word:list symbol) :=
+ forallb nullable_symb word.
+
+(** First predicate for non terminal, symbols and list of symbols, given as FSets. **)
+Definition first_nterm_set (nterm:nonterminal) :=
+ fold_left (fun acc t => TerminalSet.add t acc)
+ (first_nterm nterm) TerminalSet.empty.
+
+Definition first_symb_set (symbol:symbol) :=
+ match symbol with
+ | NT nt => first_nterm_set nt
+ | T t => TerminalSet.singleton t
+ end.
+
+Fixpoint first_word_set (word:list symbol) :=
+ match word with
+ | [] => TerminalSet.empty
+ | t::q =>
+ if nullable_symb t then
+ TerminalSet.union (first_symb_set t) (first_word_set q)
+ else
+ first_symb_set t
+ end.
+
+(** Small helper for finding the part of an item that is after the dot. **)
+Definition future_of_prod prod dot_pos : list symbol :=
+ (fix loop n lst :=
+ match n with
+ | O => lst
+ | S x => match loop x lst with [] => [] | _::q => q end
+ end)
+ dot_pos (rev' (prod_rhs_rev prod)).
+
+(** We build a fast map to store all the items of all the states. **)
+Definition items_map (_:unit): StateProdPosMap.t TerminalSet.t :=
+ fold_left (fun acc state =>
+ fold_left (fun acc item =>
+ let key := (state, prod_item item, dot_pos_item item) in
+ let data := fold_left (fun acc t => TerminalSet.add t acc)
+ (lookaheads_item item) TerminalSet.empty
+ in
+ let old :=
+ match StateProdPosMap.find key acc with
+ | Some x => x | None => TerminalSet.empty
+ end
+ in
+ StateProdPosMap.add key (TerminalSet.union data old) acc
+ ) (items_of_state state) acc
+ ) all_list (StateProdPosMap.empty TerminalSet.t).
+
+(** We need to avoid computing items_map each time we need it. To that
+ purpose, we declare a typeclass specifying that some map is equal to
+ items_map. *)
+Class IsItemsMap m := is_items_map : m = items_map ().
+
+(** Accessor. **)
+Definition find_items_map items_map state prod dot_pos : TerminalSet.t :=
+ match StateProdPosMap.find (state, prod, dot_pos) items_map with
+ | None => TerminalSet.empty
+ | Some x => x
+ end.
+
+Definition state_has_future state prod (fut:list symbol) (lookahead:terminal) :=
+ exists dot_pos:nat,
+ fut = future_of_prod prod dot_pos /\
+ TerminalSet.In lookahead (find_items_map (items_map ()) state prod dot_pos).
+
+(** Iterator over items. **)
+Definition forallb_items items_map (P:state -> production -> nat -> TerminalSet.t -> bool): bool:=
+ StateProdPosMap.fold (fun key set acc =>
+ match key with (st, p, pos) => (acc && P st p pos set)%bool end
+ ) items_map true.
+
+(** Typeclass instances for synthetizing the validator. *)
+
+Instance is_validator_subset S1 S2 :
+ IsValidator (TerminalSet.Subset S1 S2) (TerminalSet.subset S1 S2).
+Proof. intros ?. by apply TerminalSet.subset_2. Qed.
+
+(* While the specification of the validator always quantify over
+ possible lookahead tokens individually, the validator usually
+ handles lookahead sets directly instead, for better performances.
+
+ For instance, the validator for [state_has_future], which speaks
+ about one single lookahead token is a subset operation:
+*)
+Lemma is_validator_state_has_future_subset st prod pos lookahead lset im fut :
+ TerminalSet.In lookahead lset ->
+ fut = future_of_prod prod pos ->
+ IsItemsMap im ->
+ IsValidator (state_has_future st prod fut lookahead)
+ (TerminalSet.subset lset (find_items_map im st prod pos)).
+Proof.
+ intros ? -> -> HSS%TerminalSet.subset_2. exists pos. split=>//. by apply HSS.
+Qed.
+(* We do not declare this lemma as an instance, and use [Hint Extern]
+ instead, because the typeclass mechanism has trouble instantiating
+ some evars if we do not explicitely call [eassumption]. *)
+Hint Extern 2 (IsValidator (state_has_future _ _ _ _) _) =>
+ eapply is_validator_state_has_future_subset; [eassumption|eassumption || reflexivity|]
+: typeclass_instances.
+
+(* As said previously, we manipulate lookahead terminal sets instead of
+ lookahead individually. Hence, when we quantify over a lookahead set
+ in the specification, we do not do anything in the executable
+ validator.
+
+ This instance is used for [non_terminal_closed]. *)
+Instance is_validator_forall_lookahead_set lset P b:
+ (forall lookahead, TerminalSet.In lookahead lset -> IsValidator (P lookahead) b) ->
+ IsValidator (forall lookahead, TerminalSet.In lookahead lset -> P lookahead) b.
+Proof. unfold IsValidator. firstorder. Qed.
+
+
+(* Dually, we sometimes still need to explicitelly iterate over a
+ lookahead set. This is what this lemma allows.
+ Used only in [end_reduce]. *)
+Lemma is_validator_iterate_lset P b lookahead lset :
+ TerminalSet.In lookahead lset ->
+ IsValidator P (b lookahead) ->
+ IsValidator P (TerminalSet.fold (fun lookahead acc =>
+ if acc then b lookahead else false) lset true).
+Proof.
+ intros Hlset%TerminalSet.elements_1 Hval Val. apply Hval.
+ revert Val. rewrite TerminalSet.fold_1. generalize true at 1. clear -Hlset.
+ induction Hlset as [? l <-%compare_eq|? l ? IH]=> /= b' Val.
+ - destruct (b lookahead). by destruct b'. exfalso. by induction l; destruct b'.
+ - eauto.
+Qed.
+Hint Extern 100 (IsValidator _ _) =>
+ match goal with
+ | H : TerminalSet.In ?lookahead ?lset |- _ =>
+ eapply (is_validator_iterate_lset _ (fun lookahead => _) _ _ H); clear H
+ end
+: typeclass_instances.
+
+(* We often quantify over all the items of all the states of the
+ automaton. This lemma and the accompanying [Hint Resolve]
+ declaration allow generating the corresponding executable
+ validator.
+
+ Note that it turns out that, in all the uses of this pattern, the
+ first thing we do for each item is pattern-matching over the
+ future. This lemma also embbed this pattern-matching, which makes
+ it possible to get the hypothesis [fut' = future_of_prod prod (S pos)]
+ in the non-nil branch.
+
+ Moreover, note, again, that while the specification quantifies over
+ lookahead terminals individually, the code provides lookahead sets
+ instead. *)
+Lemma is_validator_forall_items P1 b1 P2 b2 im :
+ IsItemsMap im ->
+
+ (forall st prod lookahead lset pos,
+ TerminalSet.In lookahead lset ->
+ [] = future_of_prod prod pos ->
+ IsValidator (P1 st prod lookahead) (b1 st prod pos lset)) ->
+
+ (forall st prod pos lookahead lset s fut',
+ TerminalSet.In lookahead lset ->
+ fut' = future_of_prod prod (S pos) ->
+ IsValidator (P2 st prod lookahead s fut') (b2 st prod pos lset s fut')) ->
+
+ IsValidator (forall st prod fut lookahead,
+ state_has_future st prod fut lookahead ->
+ match fut with
+ | [] => P1 st prod lookahead
+ | s :: fut' => P2 st prod lookahead s fut'
+ end)
+ (forallb_items im (fun st prod pos lset =>
+ match future_of_prod prod pos with
+ | [] => b1 st prod pos lset
+ | s :: fut' => b2 st prod pos lset s fut'
+ end)).
+Proof.
+ intros -> Hval1 Hval2 Val st prod fut lookahead (pos & -> & Hlookahead).
+ rewrite /forallb_items StateProdPosMap.fold_1 in Val.
+ assert (match future_of_prod prod pos with
+ | [] => b1 st prod pos (find_items_map (items_map ()) st prod pos)
+ | s :: fut' => b2 st prod pos (find_items_map (items_map ()) st prod pos) s fut'
+ end = true).
+ - unfold find_items_map in *.
+ assert (Hfind := @StateProdPosMap.find_2 _ (items_map ()) (st, prod, pos)).
+ destruct StateProdPosMap.find as [lset|]; [|by edestruct (TerminalSet.empty_1); eauto].
+ specialize (Hfind _ eq_refl). apply StateProdPosMap.elements_1 in Hfind.
+ revert Val. generalize true at 1.
+ induction Hfind as [[? ?] l [?%compare_eq ?]|??? IH]=>?.
+ + simpl in *; subst.
+ match goal with |- _ -> ?X = true => destruct X end; [done|].
+ rewrite Bool.andb_false_r. clear. induction l as [|[[[??]?]?] l IH]=>//.
+ + apply IH.
+ - destruct future_of_prod eqn:EQ. by eapply Hval1; eauto.
+ eapply Hval2 with (pos := pos); eauto; [].
+ revert EQ. unfold future_of_prod=>-> //.
+Qed.
+(* We need a hint for expplicitely instantiating b1 and b2 with lambdas. *)
+Hint Extern 0 (IsValidator
+ (forall st prod fut lookahead,
+ state_has_future st prod fut lookahead -> _)
+ _) =>
+ eapply (is_validator_forall_items _ (fun st prod pos lset => _)
+ _ (fun st prod pos lset s fut' => _))
+ : typeclass_instances.
+
+(* Used in [start_future] only. *)
+Instance is_validator_forall_state_has_future im st prod :
+ IsItemsMap im ->
+ IsValidator
+ (forall look, state_has_future st prod (rev' (prod_rhs_rev prod)) look)
+ (let lookaheads := find_items_map im st prod 0 in
+ forallb (fun t => TerminalSet.mem t lookaheads) all_list).
+Proof.
+ move=> -> /forallb_forall Val look.
+ specialize (Val look (all_list_forall _)). exists 0. split=>//.
+ by apply TerminalSet.mem_2.
+Qed.
+
+(** * Validation for completeness **)
+
+(** The nullable predicate is a fixpoint : it is correct. **)
+Definition nullable_stable :=
+ forall p:production,
+ if nullable_word (prod_rhs_rev p) then
+ nullable_nterm (prod_lhs p) = true
+ else True.
+
+(** The first predicate is a fixpoint : it is correct. **)
+Definition first_stable:=
+ forall (p:production),
+ TerminalSet.Subset (first_word_set (rev' (prod_rhs_rev p)))
+ (first_nterm_set (prod_lhs p)).
+
+(** The initial state has all the S=>.u items, where S is the start non-terminal **)
+Definition start_future :=
+ forall (init:initstate) (p:production),
+ prod_lhs p = start_nt init ->
+ forall (t:terminal),
+ state_has_future init p (rev' (prod_rhs_rev p)) t.
+
+(** If a state contains an item of the form A->_.av[[b]], where a is a
+ terminal, then reading an a does a [Shift_act], to a state containing
+ an item of the form A->_.v[[b]]. **)
+Definition terminal_shift :=
+ forall (s1:state) prod fut lookahead,
+ state_has_future s1 prod fut lookahead ->
+ match fut with
+ | T t::q =>
+ match action_table s1 with
+ | Lookahead_act awp =>
+ match awp t with
+ | Shift_act s2 _ =>
+ state_has_future s2 prod q lookahead
+ | _ => False
+ end
+ | _ => False
+ end
+ | _ => True
+ end.
+
+(** If a state contains an item of the form A->_.[[a]], then either we do a
+ [Default_reduce_act] of the corresponding production, either a is a
+ terminal (ie. there is a lookahead terminal), and reading a does a
+ [Reduce_act] of the corresponding production. **)
+Definition end_reduce :=
+ forall (s:state) prod fut lookahead,
+ state_has_future s prod fut lookahead ->
+ match fut with
+ | [] =>
+ match action_table s with
+ | Default_reduce_act p => p = prod
+ | Lookahead_act awt =>
+ match awt lookahead with
+ | Reduce_act p => p = prod
+ | _ => False
+ end
+ end
+ | _ => True
+ end.
+
+Definition is_end_reduce items_map :=
+ forallb_items items_map (fun s prod pos lset =>
+ match future_of_prod prod pos with
+ | [] =>
+ match action_table s with
+ | Default_reduce_act p => compare_eqb p prod
+ | Lookahead_act awt =>
+ TerminalSet.fold (fun lookahead acc =>
+ match awt lookahead with
+ | Reduce_act p => (acc && compare_eqb p prod)%bool
+ | _ => false
+ end) lset true
+ end
+ | _ => true
+ end).
+
+(** If a state contains an item of the form A->_.Bv[[b]], where B is a
+ non terminal, then the goto table says we have to go to a state containing
+ an item of the form A->_.v[[b]]. **)
+Definition non_terminal_goto :=
+ forall (s1:state) prod fut lookahead,
+ state_has_future s1 prod fut lookahead ->
+ match fut with
+ | NT nt::q =>
+ match goto_table s1 nt with
+ | Some (exist _ s2 _) =>
+ state_has_future s2 prod q lookahead
+ | None => False
+ end
+ | _ => True
+ end.
+
+Definition start_goto :=
+ forall (init:initstate),
+ match goto_table init (start_nt init) with
+ | None => True
+ | Some _ => False
+ end.
+
+(** Closure property of item sets : if a state contains an item of the form
+ A->_.Bv[[b]], then for each production B->u and each terminal a of
+ first(vb), the state contains an item of the form B->_.u[[a]] **)
+Definition non_terminal_closed :=
+ forall s1 prod fut lookahead,
+ state_has_future s1 prod fut lookahead ->
+ match fut with
+ | NT nt::q =>
+ forall p, prod_lhs p = nt ->
+ (if nullable_word q then
+ state_has_future s1 p (future_of_prod p 0) lookahead
+ else True) /\
+ (forall lookahead2,
+ TerminalSet.In lookahead2 (first_word_set q) ->
+ state_has_future s1 p (future_of_prod p 0) lookahead2)
+ | _ => True
+ end.
+
+(** The automaton is complete **)
+Definition complete :=
+ nullable_stable /\ first_stable /\ start_future /\ terminal_shift
+ /\ end_reduce /\ non_terminal_goto /\ start_goto /\ non_terminal_closed.
+
+Derive is_complete_0
+SuchThat (forall im, IsItemsMap im -> IsValidator complete (is_complete_0 im))
+As complete_0_is_validator.
+Proof. intros im. subst is_complete_0. instantiate (1:=fun im => _). apply _. Qed.
+
+Definition is_complete (_:unit) := is_complete_0 (items_map ()).
+Lemma complete_is_validator : IsValidator complete (is_complete ()).
+Proof. by apply complete_0_is_validator. Qed.
+
+End Make.
diff --git a/MenhirLib/Validator_safe.v b/MenhirLib/Validator_safe.v
new file mode 100644
index 00000000..628d2009
--- /dev/null
+++ b/MenhirLib/Validator_safe.v
@@ -0,0 +1,234 @@
+(****************************************************************************)
+(* *)
+(* Menhir *)
+(* *)
+(* Jacques-Henri Jourdan, CNRS, LRI, Université Paris Sud *)
+(* *)
+(* Copyright Inria. All rights reserved. This file is distributed under *)
+(* the terms of the GNU Lesser General Public License as published by the *)
+(* Free Software Foundation, either version 3 of the License, or (at your *)
+(* option) any later version, as described in the file LICENSE. *)
+(* *)
+(****************************************************************************)
+
+From Coq Require Import List Syntax Derive.
+From Coq.ssr Require Import ssreflect.
+Require Automaton.
+Require Import Alphabet Validator_classes.
+
+Module Make(Import A:Automaton.T).
+
+(** The singleton predicate for states **)
+Definition singleton_state_pred (state:state) :=
+ (fun state' => match compare state state' with Eq => true |_ => false end).
+
+(** [past_state_of_non_init_state], extended for all states. **)
+Definition past_state_of_state (state:state) :=
+ match state with
+ | Init _ => []
+ | Ninit nis => past_state_of_non_init_state nis
+ end.
+
+(** Concatenations of last and past **)
+Definition head_symbs_of_state (state:state) :=
+ match state with
+ | Init _ => []
+ | Ninit s =>
+ last_symb_of_non_init_state s::past_symb_of_non_init_state s
+ end.
+Definition head_states_of_state (state:state) :=
+ singleton_state_pred state::past_state_of_state state.
+
+(** * Validation for correctness **)
+
+(** Prefix predicate between two lists of symbols. **)
+Inductive prefix: list symbol -> list symbol -> Prop :=
+| prefix_nil: forall l, prefix [] l
+| prefix_cons: forall l1 l2 x, prefix l1 l2 -> prefix (x::l1) (x::l2).
+
+(** [prefix] is transitive **)
+Lemma prefix_trans:
+ forall (l1 l2 l3:list symbol), prefix l1 l2 -> prefix l2 l3 -> prefix l1 l3.
+Proof.
+ intros l1 l2 l3 H1 H2. revert l3 H2.
+ induction H1; [now constructor|]. inversion 1. subst. constructor. eauto.
+Qed.
+
+Fixpoint is_prefix (l1 l2:list symbol) :=
+ match l1, l2 with
+ | [], _ => true
+ | t1::q1, t2::q2 => (compare_eqb t1 t2 && is_prefix q1 q2)%bool
+ | _::_, [] => false
+ end.
+
+Instance prefix_is_validator l1 l2 : IsValidator (prefix l1 l2) (is_prefix l1 l2).
+Proof.
+ revert l2. induction l1 as [|x1 l1 IH]=>l2 Hpref.
+ - constructor.
+ - destruct l2 as [|x2 l2]=>//.
+ move: Hpref=> /andb_prop [/compare_eqb_iff -> /IH ?]. by constructor.
+Qed.
+
+(** If we shift, then the known top symbols of the destination state is
+ a prefix of the known top symbols of the source state, with the new
+ symbol added. **)
+Definition shift_head_symbs :=
+ forall s,
+ match action_table s with
+ | Lookahead_act awp => forall t,
+ match awp t with
+ | Shift_act s2 _ =>
+ prefix (past_symb_of_non_init_state s2) (head_symbs_of_state s)
+ | _ => True
+ end
+ | _ => True
+ end.
+
+(** When a goto happens, then the known top symbols of the destination state
+ is a prefix of the known top symbols of the source state, with the new
+ symbol added. **)
+Definition goto_head_symbs :=
+ forall s nt,
+ match goto_table s nt with
+ | Some (exist _ s2 _) =>
+ prefix (past_symb_of_non_init_state s2) (head_symbs_of_state s)
+ | None => True
+ end.
+
+(** We have to say the same kind of checks for the assumptions about the
+ states stack. However, theses assumptions are predicates. So we define
+ a notion of "prefix" over predicates lists, that means, basically, that
+ an assumption entails another **)
+Inductive prefix_pred: list (state->bool) -> list (state->bool) -> Prop :=
+ | prefix_pred_nil: forall l, prefix_pred [] l
+ | prefix_pred_cons: forall l1 l2 f1 f2,
+ (forall x, implb (f2 x) (f1 x) = true) ->
+ prefix_pred l1 l2 -> prefix_pred (f1::l1) (f2::l2).
+
+(** [prefix_pred] is transitive **)
+Lemma prefix_pred_trans:
+ forall (l1 l2 l3:list (state->bool)),
+ prefix_pred l1 l2 -> prefix_pred l2 l3 -> prefix_pred l1 l3.
+Proof.
+ intros l1 l2 l3 H1 H2. revert l3 H2.
+ induction H1 as [|l1 l2 f1 f2 Hf2f1]; [now constructor|].
+ intros l3. inversion 1 as [|??? f3 Hf3f2]. subst. constructor; [|now eauto].
+ intros x. specialize (Hf3f2 x). specialize (Hf2f1 x).
+ repeat destruct (_ x); auto.
+Qed.
+
+Fixpoint is_prefix_pred (l1 l2:list (state->bool)) :=
+ match l1, l2 with
+ | [], _ => true
+ | f1::q1, f2::q2 =>
+ (forallb (fun x => implb (f2 x) (f1 x)) all_list
+ && is_prefix_pred q1 q2)%bool
+ | _::_, [] => false
+ end.
+
+Instance prefix_pred_is_validator l1 l2 :
+ IsValidator (prefix_pred l1 l2) (is_prefix_pred l1 l2).
+Proof.
+ revert l2. induction l1 as [|x1 l1 IH]=>l2 Hpref.
+ - constructor.
+ - destruct l2 as [|x2 l2]=>//.
+ move: Hpref=> /andb_prop [/forallb_forall ? /IH ?].
+ constructor; auto using all_list_forall.
+Qed.
+
+(** The assumptions about state stack is conserved when we shift **)
+Definition shift_past_state :=
+ forall s,
+ match action_table s with
+ | Lookahead_act awp => forall t,
+ match awp t with
+ | Shift_act s2 _ =>
+ prefix_pred (past_state_of_non_init_state s2)
+ (head_states_of_state s)
+ | _ => True
+ end
+ | _ => True
+ end.
+
+(** The assumptions about state stack is conserved when we do a goto **)
+Definition goto_past_state :=
+ forall s nt,
+ match goto_table s nt with
+ | Some (exist _ s2 _) =>
+ prefix_pred (past_state_of_non_init_state s2)
+ (head_states_of_state s)
+ | None => True
+ end.
+
+(** What states are possible after having popped these symbols from the
+ stack, given the annotation of the current state ? **)
+Inductive state_valid_after_pop (s:state):
+ list symbol -> list (state -> bool) -> Prop :=
+ | state_valid_after_pop_nil1:
+ forall p pl, p s = true -> state_valid_after_pop s [] (p::pl)
+ | state_valid_after_pop_nil2:
+ forall sl, state_valid_after_pop s sl []
+ | state_valid_after_pop_cons:
+ forall st sq p pl, state_valid_after_pop s sq pl ->
+ state_valid_after_pop s (st::sq) (p::pl).
+
+Fixpoint is_state_valid_after_pop (state:state) (to_pop:list symbol) annot :=
+ match annot, to_pop with
+ | [], _ => true
+ | p::_, [] => p state
+ | p::pl, s::sl => is_state_valid_after_pop state sl pl
+ end.
+
+Instance impl_is_state_valid_after_pop_is_validator state sl pl P b :
+ IsValidator P b ->
+ IsValidator (state_valid_after_pop state sl pl -> P)
+ (if is_state_valid_after_pop state sl pl then b else true).
+Proof.
+ destruct (is_state_valid_after_pop _ sl pl) eqn:EQ.
+ - intros ???. by eapply is_validator.
+ - intros _ _ Hsvap. exfalso. induction Hsvap=>//; [simpl in EQ; congruence|].
+ by destruct sl.
+Qed.
+
+(** A state is valid for reducing a production when :
+ - The assumptions on the state are such that we will find the right hand
+ side of the production on the stack.
+ - We will be able to do a goto after having popped the right hand side.
+**)
+Definition valid_for_reduce (state:state) prod :=
+ prefix (prod_rhs_rev prod) (head_symbs_of_state state) /\
+ forall state_new,
+ state_valid_after_pop state_new
+ (prod_rhs_rev prod) (head_states_of_state state) ->
+ match goto_table state_new (prod_lhs prod) with
+ | None =>
+ match state_new with
+ | Init i => prod_lhs prod = start_nt i
+ | Ninit _ => False
+ end
+ | _ => True
+ end.
+
+(** All the states that does a reduce are valid for reduction **)
+Definition reduce_ok :=
+ forall s,
+ match action_table s with
+ | Lookahead_act awp =>
+ forall t, match awp t with
+ | Reduce_act p => valid_for_reduce s p
+ | _ => True
+ end
+ | Default_reduce_act p => valid_for_reduce s p
+ end.
+
+(** The automaton is safe **)
+Definition safe :=
+ shift_head_symbs /\ goto_head_symbs /\ shift_past_state /\
+ goto_past_state /\ reduce_ok.
+
+Derive is_safe
+SuchThat (IsValidator safe (is_safe ()))
+As safe_is_validator.
+Proof. subst is_safe. instantiate (1:=fun _ => _). apply _. Qed.
+
+End Make.
diff --git a/arm/Archi.v b/arm/Archi.v
index 39a424ec..16d6c71d 100644
--- a/arm/Archi.v
+++ b/arm/Archi.v
@@ -16,7 +16,7 @@
(** Architecture-dependent parameters for ARM *)
-Require Import ZArith.
+Require Import ZArith List.
(*From Flocq*)
Require Import Binary Bits.
@@ -34,30 +34,57 @@ Proof.
unfold splitlong, ptr64; congruence.
Qed.
-Definition default_nan_64 : { x : binary64 | is_nan _ _ x = true } :=
- exist _ (B754_nan 53 1024 false (iter_nat 51 _ xO xH) (eq_refl true)) (eq_refl true).
+Definition default_nan_64 := (false, iter_nat 51 _ xO xH).
+Definition default_nan_32 := (false, iter_nat 22 _ xO xH).
+
+(** Choose the first signaling NaN, if any;
+ otherwise choose the first NaN;
+ otherwise use default. *)
+
+Definition choose_nan (is_signaling: positive -> bool)
+ (default: bool * positive)
+ (l0: list (bool * positive)) : bool * positive :=
+ let fix choose_snan (l1: list (bool * positive)) :=
+ match l1 with
+ | nil =>
+ match l0 with nil => default | n :: _ => n end
+ | ((s, p) as n) :: l1 =>
+ if is_signaling p then n else choose_snan l1
+ end
+ in choose_snan l0.
+
+Lemma choose_nan_idem: forall is_signaling default n,
+ choose_nan is_signaling default (n :: n :: nil) =
+ choose_nan is_signaling default (n :: nil).
+Proof.
+ intros. destruct n as [s p]; unfold choose_nan; simpl.
+ destruct (is_signaling p); auto.
+Qed.
+
+Definition choose_nan_64 :=
+ choose_nan (fun p => negb (Pos.testbit p 51)) default_nan_64.
+
+Definition choose_nan_32 :=
+ choose_nan (fun p => negb (Pos.testbit p 22)) default_nan_32.
-Definition choose_binop_pl_64 (pl1 pl2 : positive) :=
- (** Choose second NaN if pl2 is sNaN but pl1 is qNan.
- In all other cases, choose first NaN *)
- (Pos.testbit pl1 51 && negb (Pos.testbit pl2 51))%bool.
+Lemma choose_nan_64_idem: forall n,
+ choose_nan_64 (n :: n :: nil) = choose_nan_64 (n :: nil).
+Proof. intros; apply choose_nan_idem. Qed.
-Definition default_nan_32 : { x : binary32 | is_nan _ _ x = true } :=
- exist _ (B754_nan 24 128 false (iter_nat 22 _ xO xH) (eq_refl true)) (eq_refl true).
+Lemma choose_nan_32_idem: forall n,
+ choose_nan_32 (n :: n :: nil) = choose_nan_32 (n :: nil).
+Proof. intros; apply choose_nan_idem. Qed.
-Definition choose_binop_pl_32 (pl1 pl2 : positive) :=
- (** Choose second NaN if pl2 is sNaN but pl1 is qNan.
- In all other cases, choose first NaN *)
- (Pos.testbit pl1 22 && negb (Pos.testbit pl2 22))%bool.
+Definition fma_order {A: Type} (x y z: A) := (z, x, y).
-Definition fpu_returns_default_qNaN := false.
+Definition fma_invalid_mul_is_nan := true.
Definition float_of_single_preserves_sNaN := false.
Global Opaque ptr64 big_endian splitlong
- default_nan_64 choose_binop_pl_64
- default_nan_32 choose_binop_pl_32
- fpu_returns_default_qNaN
+ default_nan_64 choose_nan_64
+ default_nan_32 choose_nan_32
+ fma_order fma_invalid_mul_is_nan
float_of_single_preserves_sNaN.
(** Which ABI to use: either the standard ARM EABI with floats passed
diff --git a/arm/Asmexpand.ml b/arm/Asmexpand.ml
index d9424d11..a4ec0c5d 100644
--- a/arm/Asmexpand.ml
+++ b/arm/Asmexpand.ml
@@ -304,6 +304,11 @@ let expand_builtin_va_start r =
let expand_builtin_inline name args res =
match name, args, res with
(* Integer arithmetic *)
+ | "__builtin_bswap64" , [BA_splitlong(BA(IR ah), BA(IR al))],
+ BR_splitlong(BR(IR rh), BR(IR rl)) ->
+ expand_int64_arith (rl = al) rl (fun rl ->
+ emit (Prev (rl, ah));
+ emit (Prev (rh, al)))
| ("__builtin_bswap" | "__builtin_bswap32"), [BA(IR a1)], BR(IR res) ->
emit (Prev (res, a1))
| "__builtin_bswap16", [BA(IR a1)], BR(IR res) ->
diff --git a/cparser/Builtins.mli b/arm/Builtins1.v
index 7f9d78a9..f6e643d2 100644
--- a/cparser/Builtins.mli
+++ b/arm/Builtins1.v
@@ -2,7 +2,7 @@
(* *)
(* The Compcert verified compiler *)
(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
+(* Xavier Leroy, Collège de France and Inria Paris *)
(* *)
(* Copyright Institut National de Recherche en Informatique et en *)
(* Automatique. All rights reserved. This file is distributed *)
@@ -13,13 +13,21 @@
(* *)
(* *********************************************************************)
-val environment: unit -> Env.t
-val identifiers: unit -> C.ident list
-val declarations: unit -> C.globdecl list
+(** Platform-specific built-in functions *)
-type t = {
- typedefs: (string * C.typ) list;
- functions: (string * (C.typ * C.typ list * bool)) list
-}
+Require Import String Coqlib.
+Require Import AST Integers Floats Values.
+Require Import Builtins0.
-val set: t -> unit
+Inductive platform_builtin : Type := .
+
+Local Open Scope string_scope.
+
+Definition platform_builtin_table : list (string * platform_builtin) :=
+ nil.
+
+Definition platform_builtin_sig (b: platform_builtin) : signature :=
+ match b with end.
+
+Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (proj_sig_res (platform_builtin_sig b)) :=
+ match b with end.
diff --git a/arm/CBuiltins.ml b/arm/CBuiltins.ml
index ec4f4aaa..d6a1ea35 100644
--- a/arm/CBuiltins.ml
+++ b/arm/CBuiltins.ml
@@ -18,10 +18,10 @@
open C
let builtins = {
- Builtins.typedefs = [
+ builtin_typedefs = [
"__builtin_va_list", TPtr(TVoid [], [])
];
- Builtins.functions = [
+ builtin_functions = [
(* Integer arithmetic *)
"__builtin_clz",
(TInt(IInt, []), [TInt(IUInt, [])], false);
diff --git a/arm/ConstpropOp.vp b/arm/ConstpropOp.vp
index d62240ef..8555d3aa 100644
--- a/arm/ConstpropOp.vp
+++ b/arm/ConstpropOp.vp
@@ -20,7 +20,7 @@ Require Import Integers.
Require Import Floats.
Require Import Op.
Require Import Registers.
-Require Import ValueDomain.
+Require Import ValueDomain ValueAOp.
(** * Converting known values to constants *)
@@ -131,6 +131,15 @@ Nondetfunction make_cmp (c: condition) (args: list reg) (vl: list aval) :=
make_cmp_base c args vl
end.
+Definition make_select (c: condition) (ty: typ)
+ (r1 r2: reg) (args: list reg) (vl: list aval) :=
+ match resolve_branch (eval_static_condition c vl) with
+ | Some b => (Omove, (if b then r1 else r2) :: nil)
+ | None =>
+ let (c', args') := cond_strength_reduction c args vl in
+ (Osel c' ty, r1 :: r2 :: args')
+ end.
+
Definition make_addimm (n: int) (r: reg) :=
if Int.eq n Int.zero
then (Omove, r :: nil)
@@ -284,6 +293,7 @@ Nondetfunction op_strength_reduction
| Oshr, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrimm n2 r1 r2
| Oshru, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shruimm n2 r1 r2
| Ocmp c, args, vl => make_cmp c args vl
+ | Osel c ty, r1 :: r2 :: args, v1 :: v2 :: vl => make_select c ty r1 r2 args vl
| Omulf, r1 :: r2 :: nil, v1 :: F n2 :: nil => make_mulfimm n2 r1 r1 r2
| Omulf, r1 :: r2 :: nil, F n1 :: v2 :: nil => make_mulfimm n1 r2 r1 r2
| Omulfs, r1 :: r2 :: nil, v1 :: FS n2 :: nil => make_mulfsimm n2 r1 r1 r2
diff --git a/arm/ConstpropOpproof.v b/arm/ConstpropOpproof.v
index 079ba2be..a4f5c29c 100644
--- a/arm/ConstpropOpproof.v
+++ b/arm/ConstpropOpproof.v
@@ -24,7 +24,7 @@ Require Import Events.
Require Import Op.
Require Import Registers.
Require Import RTL.
-Require Import ValueDomain.
+Require Import ValueDomain ValueAOp ValueAnalysis.
Require Import ConstpropOp.
Local Transparent Archi.ptr64.
@@ -234,6 +234,28 @@ Proof.
- apply make_cmp_base_correct; auto.
Qed.
+Lemma make_select_correct:
+ forall c ty r1 r2 args vl,
+ vl = map (fun r => AE.get r ae) args ->
+ let (op', args') := make_select c ty r1 r2 args vl in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op' rs##args' m = Some v
+ /\ Val.lessdef (Val.select (eval_condition c rs##args m) rs#r1 rs#r2 ty) v.
+Proof.
+ unfold make_select; intros.
+ destruct (resolve_branch (eval_static_condition c vl)) as [b|] eqn:RB.
+- exists (if b then rs#r1 else rs#r2); split.
++ simpl. destruct b; auto.
++ destruct (eval_condition c rs##args m) as [b'|] eqn:EC; simpl; auto.
+ assert (b = b').
+ { eapply resolve_branch_sound; eauto.
+ rewrite <- EC. apply eval_static_condition_sound with bc.
+ subst vl. exact (aregs_sound _ _ _ args MATCH). }
+ subst b'. apply Val.lessdef_normalize.
+- generalize (cond_strength_reduction_correct c args vl H).
+ destruct (cond_strength_reduction c args vl) as [cond' args']; intros EQ.
+ econstructor; split. simpl; eauto. rewrite EQ; auto.
+Qed.
+
Lemma make_addimm_correct:
forall n r,
let (op, args) := make_addimm n r in
@@ -616,6 +638,8 @@ Proof.
InvApproxRegs; SimplVM. inv H0. apply make_shruimm_correct; auto.
(* cmp *)
inv H0. apply make_cmp_correct; auto.
+(* select *)
+ inv H0. apply make_select_correct; congruence.
(* mulf *)
InvApproxRegs; SimplVM; inv H0. rewrite <- H2. apply make_mulfimm_correct; auto.
InvApproxRegs; SimplVM; inv H0. fold (Val.mulf (Vfloat n1) rs#r2).
diff --git a/arm/SelectOp.vp b/arm/SelectOp.vp
index e3ef3eaf..5506157c 100644
--- a/arm/SelectOp.vp
+++ b/arm/SelectOp.vp
@@ -38,12 +38,8 @@
Require Import Coqlib.
Require Import Compopts.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Op.
-Require Import OpHelpers.
-Require Import CminorSel.
+Require Import AST Integers Floats Builtins.
+Require Import Op OpHelpers CminorSel.
Local Open Scope cminorsel_scope.
@@ -526,3 +522,8 @@ Definition divf_base (e1: expr) (e2: expr) :=
Definition divfs_base (e1: expr) (e2: expr) :=
Eop Odivfs (e1 ::: e2 ::: Enil).
+
+(** Platform-specific known builtins *)
+
+Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr :=
+ None.
diff --git a/arm/SelectOpproof.v b/arm/SelectOpproof.v
index f7dd8dd6..75d32ea3 100644
--- a/arm/SelectOpproof.v
+++ b/arm/SelectOpproof.v
@@ -13,16 +13,9 @@
(** Correctness of instruction selection for operators *)
Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Cminor.
-Require Import Op.
-Require Import CminorSel.
+Require Import AST Integers Floats.
+Require Import Values Memory Builtins Globalenvs.
+Require Import Cminor Op CminorSel.
Require Import SelectOp.
Require Import OpHelpers OpHelpersproof.
@@ -912,7 +905,6 @@ Proof.
- constructor; auto.
Qed.
-
(* floating-point division without HELPERS *)
Theorem eval_divf_base:
forall le a b x y,
@@ -933,4 +925,17 @@ Proof.
intros; unfold divfs_base.
TrivialExists.
Qed.
+
+(** Platform-specific known builtins *)
+
+Theorem eval_platform_builtin:
+ forall bf al a vl v le,
+ platform_builtin bf al = Some a ->
+ eval_exprlist ge sp e m le al vl ->
+ platform_builtin_sem bf vl = Some v ->
+ exists v', eval_expr ge sp e m le a v' /\ Val.lessdef v v'.
+Proof.
+ intros. discriminate.
+Qed.
+
End CMCONSTR.
diff --git a/arm/TargetPrinter.ml b/arm/TargetPrinter.ml
index 3a0814e1..03e06a65 100644
--- a/arm/TargetPrinter.ml
+++ b/arm/TargetPrinter.ml
@@ -113,9 +113,9 @@ struct
let freg_single oc r = output_string oc (single_float_reg_name r)
let freg_param_single oc r = output_string oc (single_param_reg_name r)
- let preg oc = function
+ let preg_asm oc ty = function
| IR r -> ireg oc r
- | FR r -> freg oc r
+ | FR r -> if ty = Tsingle then freg_single oc r else freg oc r
| _ -> assert false
(* In Thumb2 mode, some arithmetic instructions have shorter encodings
@@ -480,7 +480,7 @@ struct
(P.to_int kind) (extern_atom txt) args
| EF_inline_asm(txt, sg, clob) ->
fprintf oc "%s begin inline assembly\n\t" comment;
- print_inline_asm preg oc (camlstring_of_coqstring txt) sg args res;
+ print_inline_asm preg_asm oc (camlstring_of_coqstring txt) sg args res;
fprintf oc "%s end inline assembly\n" comment
| _ ->
assert false
diff --git a/backend/CSE.v b/backend/CSE.v
index aaaf492c..9da30f50 100644
--- a/backend/CSE.v
+++ b/backend/CSE.v
@@ -14,7 +14,7 @@
proceeds by value numbering over extended basic blocks. *)
Require Import Coqlib Maps Errors Integers Floats Lattice Kildall.
-Require Import AST Linking.
+Require Import AST Linking Builtins.
Require Import Values Memory.
Require Import Op Registers RTL.
Require Import ValueDomain ValueAnalysis CSEdomain CombineOp.
@@ -444,10 +444,10 @@ Module Solver := BBlock_solver(Numbering).
([EF_external], [EF_runtime], [EF_malloc], [EF_free]).
- Forget equations involving loads but keep equations over registers.
This is appropriate for builtins that can modify memory,
- e.g. volatile stores, or [EF_builtin]
+ e.g. volatile stores, or [EF_builtin] for unknown builtin functions.
- Keep all equations, taking advantage of the fact that neither memory
- nor registers are modified. This is appropriate for annotations
- and for volatile loads.
+ nor registers are modified. This is appropriate for annotations,
+ volatile loads, and known builtin functions.
*)
Definition transfer (f: function) (approx: PMap.t VA.t) (pc: node) (before: numbering) :=
@@ -473,8 +473,13 @@ Definition transfer (f: function) (approx: PMap.t VA.t) (pc: node) (before: numb
match ef with
| EF_external _ _ | EF_runtime _ _ | EF_malloc | EF_free | EF_inline_asm _ _ _ =>
empty_numbering
- | EF_builtin _ _ | EF_vstore _ =>
+ | EF_vstore _ =>
set_res_unknown (kill_all_loads before) res
+ | EF_builtin name sg =>
+ match lookup_builtin_function name sg with
+ | Some bf => set_res_unknown before res
+ | None => set_res_unknown (kill_all_loads before) res
+ end
| EF_memcpy sz al =>
match args with
| dst :: src :: nil =>
diff --git a/backend/CSEproof.v b/backend/CSEproof.v
index a60c316b..03c7ecfc 100644
--- a/backend/CSEproof.v
+++ b/backend/CSEproof.v
@@ -14,7 +14,7 @@
Require Import Coqlib Maps Errors Integers Floats Lattice Kildall.
Require Import AST Linking.
-Require Import Values Memory Events Globalenvs Smallstep.
+Require Import Values Memory Builtins Events Globalenvs Smallstep.
Require Import Op Registers RTL.
Require Import ValueDomain ValueAOp ValueAnalysis.
Require Import CSEdomain CombineOp CombineOpproof CSE.
@@ -1129,7 +1129,9 @@ Proof.
{ exists valu. apply set_res_unknown_holds. eapply kill_all_loads_hold; eauto. }
destruct ef.
+ apply CASE1.
- + apply CASE3.
+ + destruct (lookup_builtin_function name sg) as [bf|] eqn:LK.
+ ++ apply CASE2. simpl in H1; red in H1; rewrite LK in H1; inv H1. auto.
+ ++ apply CASE3.
+ apply CASE1.
+ apply CASE2; inv H1; auto.
+ apply CASE3.
diff --git a/backend/Constprop.v b/backend/Constprop.v
index b68f4cae..cf1a9171 100644
--- a/backend/Constprop.v
+++ b/backend/Constprop.v
@@ -15,7 +15,7 @@
and the corresponding code rewriting. *)
Require Import Coqlib Maps Integers Floats Lattice Kildall.
-Require Import AST Linking.
+Require Import AST Linking Builtins.
Require Compopts Machregs.
Require Import Op Registers RTL.
Require Import Liveness ValueDomain ValueAOp ValueAnalysis.
@@ -139,6 +139,30 @@ Definition builtin_strength_reduction
| _ => builtin_args_strength_reduction ae al (Machregs.builtin_constraints ef)
end.
+(*
+Definition transf_builtin
+ (ae: AE.t) (am: amem) (rm: romem)
+ (ef: external_function)
+ (args: list (builtin_arg reg)) (res: builtin_res reg) (s: node) :=
+ let dfl := Ibuiltin ef (builtin_strength_reduction ae ef args) res s in
+ match ef, res with
+ | EF_builtin name sg, BR rd =>
+ match lookup_builtin_function name sg with
+ | Some bf =>
+ match eval_static_builtin_function ae am rm bf args with
+ | Some a =>
+ match const_for_result a with
+ | Some cop => Iop cop nil rd s
+ | None => dfl
+ end
+ | None => dfl
+ end
+ | None => dfl
+ end
+ | _, _ => dfl
+ end.
+*)
+
Definition transf_instr (f: function) (an: PMap.t VA.t) (rm: romem)
(pc: node) (instr: instruction) :=
match an!!pc with
@@ -176,7 +200,23 @@ Definition transf_instr (f: function) (an: PMap.t VA.t) (rm: romem)
| Itailcall sig ros args =>
Itailcall sig (transf_ros ae ros) args
| Ibuiltin ef args res s =>
- Ibuiltin ef (builtin_strength_reduction ae ef args) res s
+ let dfl := Ibuiltin ef (builtin_strength_reduction ae ef args) res s in
+ match ef, res with
+ | EF_builtin name sg, BR rd =>
+ match lookup_builtin_function name sg with
+ | Some bf =>
+ match eval_static_builtin_function ae am rm bf args with
+ | Some a =>
+ match const_for_result a with
+ | Some cop => Iop cop nil rd s
+ | None => dfl
+ end
+ | None => dfl
+ end
+ | None => dfl
+ end
+ | _, _ => dfl
+ end
| Icond cond args s1 s2 =>
let aargs := aregs ae args in
match resolve_branch (eval_static_condition cond aargs) with
diff --git a/backend/Constpropproof.v b/backend/Constpropproof.v
index e28519ca..a5d08a0f 100644
--- a/backend/Constpropproof.v
+++ b/backend/Constpropproof.v
@@ -14,7 +14,7 @@
Require Import Coqlib Maps Integers Floats Lattice Kildall.
Require Import AST Linking.
-Require Import Values Events Memory Globalenvs Smallstep.
+Require Import Values Builtins Events Memory Globalenvs Smallstep.
Require Compopts Machregs.
Require Import Op Registers RTL.
Require Import Liveness ValueDomain ValueAOp ValueAnalysis.
@@ -474,19 +474,41 @@ Proof.
- (* Ibuiltin *)
rename pc'0 into pc. TransfInstr; intros.
Opaque builtin_strength_reduction.
- exploit builtin_strength_reduction_correct; eauto. intros (vargs' & P & Q).
- exploit (@eval_builtin_args_lessdef _ ge (fun r => rs#r) (fun r => rs'#r)).
+ set (dfl := Ibuiltin ef (builtin_strength_reduction ae ef args) res pc') in *.
+ set (rm := romem_for cu) in *.
+ assert (DFL: (fn_code (transf_function rm f))!pc = Some dfl ->
+ exists (n2 : nat) (s2' : state),
+ step tge
+ (State s' (transf_function rm f) (Vptr sp0 Ptrofs.zero) pc rs' m'0) t s2' /\
+ match_states n2
+ (State s f (Vptr sp0 Ptrofs.zero) pc' (regmap_setres res vres rs) m') s2').
+ {
+ exploit builtin_strength_reduction_correct; eauto. intros (vargs' & P & Q).
+ exploit (@eval_builtin_args_lessdef _ ge (fun r => rs#r) (fun r => rs'#r)).
apply REGS. eauto. eexact P.
- intros (vargs'' & U & V).
- exploit external_call_mem_extends; eauto.
- intros [v' [m2' [A [B [C D]]]]].
+ intros (vargs'' & U & V).
+ exploit external_call_mem_extends; eauto.
+ intros (v' & m2' & A & B & C & D).
+ econstructor; econstructor; split.
+ eapply exec_Ibuiltin; eauto.
+ eapply eval_builtin_args_preserved. eexact symbols_preserved. eauto.
+ eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ eapply match_states_succ; eauto.
+ apply set_res_lessdef; auto.
+ }
+ destruct ef; auto.
+ destruct res; auto.
+ destruct (lookup_builtin_function name sg) as [bf|] eqn:LK; auto.
+ destruct (eval_static_builtin_function ae am rm bf args) as [a|] eqn:ES; auto.
+ destruct (const_for_result a) as [cop|] eqn:CR; auto.
+ clear DFL. simpl in H1; red in H1; rewrite LK in H1; inv H1.
+ exploit const_for_result_correct; eauto.
+ eapply eval_static_builtin_function_sound; eauto.
+ intros (v' & A & B).
left; econstructor; econstructor; split.
- eapply exec_Ibuiltin; eauto.
- eapply eval_builtin_args_preserved. eexact symbols_preserved. eauto.
- eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ eapply exec_Iop; eauto.
eapply match_states_succ; eauto.
- apply set_res_lessdef; auto.
-
+ apply set_reg_lessdef; auto.
- (* Icond, preserved *)
rename pc'0 into pc. TransfInstr.
set (ac := eval_static_condition cond (aregs ae args)).
diff --git a/backend/Conventions.v b/backend/Conventions.v
index 989bfa05..6025c6b4 100644
--- a/backend/Conventions.v
+++ b/backend/Conventions.v
@@ -128,8 +128,6 @@ Definition callee_save_loc (l: loc) :=
| S sl ofs ty => sl <> Outgoing
end.
-Hint Unfold callee_save_loc.
-
Definition agree_callee_save (ls1 ls2: Locmap.t) : Prop :=
forall l, callee_save_loc l -> ls1 l = ls2 l.
diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml
index 46d5c3f1..902724e0 100644
--- a/backend/Linearizeaux.ml
+++ b/backend/Linearizeaux.ml
@@ -106,7 +106,7 @@ let flatten_blocks blks =
let cmp_minpc (mpc1, _) (mpc2, _) =
if mpc1 = mpc2 then 0 else if mpc1 > mpc2 then -1 else 1
in
- List.flatten (List.map Pervasives.snd (List.sort cmp_minpc blks))
+ List.flatten (List.map snd (List.sort cmp_minpc blks))
(* Build the enumeration *)
diff --git a/backend/OpHelpersproof.v b/backend/OpHelpersproof.v
index 63040c5f..08da8a36 100644
--- a/backend/OpHelpersproof.v
+++ b/backend/OpHelpersproof.v
@@ -75,4 +75,4 @@ Definition helper_functions_declared {F V: Type} (p: AST.program (AST.fundef F)
/\ helper_declared p i32_umod "__compcert_i32_umod" sig_ii_i
/\ helper_declared p f32_div "__compcert_f32_div" sig_ss_s
/\ helper_declared p f64_div "__compcert_f64_div" sig_ff_f
-.
+. \ No newline at end of file
diff --git a/backend/PrintAsm.ml b/backend/PrintAsm.ml
index dd428808..155f5e55 100644
--- a/backend/PrintAsm.ml
+++ b/backend/PrintAsm.ml
@@ -13,7 +13,6 @@
open AST
open Camlcoq
-open DwarfPrinter
open PrintAsmaux
open Printf
open Sections
@@ -177,7 +176,7 @@ module Printer(Target:TARGET) =
let address = Target.address
end
- module DebugPrinter = DwarfPrinter (DwarfTarget)
+ module DebugPrinter = DwarfPrinter.DwarfPrinter (DwarfTarget)
end
let print_program oc p =
diff --git a/backend/PrintAsmaux.ml b/backend/PrintAsmaux.ml
index 7e075f04..8652b2c5 100644
--- a/backend/PrintAsmaux.ml
+++ b/backend/PrintAsmaux.ml
@@ -245,14 +245,15 @@ let print_debug_info comment print_line preg_string sp_name oc kind txt args =
(** Inline assembly *)
-let print_asm_argument print_preg oc modifier = function
- | BA r -> print_preg oc r
+let print_asm_argument print_preg oc modifier typ = function
+ | BA r -> print_preg oc typ r
| BA_splitlong(BA hi, BA lo) ->
begin match modifier with
- | "R" -> print_preg oc hi
- | "Q" -> print_preg oc lo
- | _ -> fprintf oc "%a:%a" print_preg hi print_preg lo
- (* Probably not what was intended *)
+ | "R" -> print_preg oc Tint hi
+ | "Q" -> print_preg oc Tint lo
+ | _ -> print_preg oc Tint hi; fprintf oc ":"; print_preg oc Tint lo
+ (* This case (printing a split long in full) should never
+ happen because of the checks done in ExtendedAsm.ml *)
end
| _ -> failwith "bad asm argument"
@@ -265,8 +266,10 @@ let re_asm_param_1 = Str.regexp "%%\\|%[QR]?[0-9]+"
let re_asm_param_2 = Str.regexp "%\\([QR]?\\)\\([0-9]+\\)"
let print_inline_asm print_preg oc txt sg args res =
- let operands =
- if sg.sig_res = None then args else builtin_arg_of_res res :: args in
+ let (operands, ty_operands) =
+ match sg.sig_res with
+ | None -> (args, sg.sig_args)
+ | Some tres -> (builtin_arg_of_res res :: args, tres :: sg.sig_args) in
let print_fragment = function
| Str.Text s ->
output_string oc s
@@ -277,7 +280,9 @@ let print_inline_asm print_preg oc txt sg args res =
let modifier = Str.matched_group 1 s
and number = int_of_string (Str.matched_group 2 s) in
try
- print_asm_argument print_preg oc modifier (List.nth operands number)
+ print_asm_argument print_preg oc modifier
+ (List.nth ty_operands number)
+ (List.nth operands number)
with Failure _ ->
fprintf oc "<bad parameter %s>" s in
List.iter print_fragment (Str.full_split re_asm_param_1 txt);
diff --git a/backend/PrintCminor.ml b/backend/PrintCminor.ml
index f68c1267..8c255a65 100644
--- a/backend/PrintCminor.ml
+++ b/backend/PrintCminor.ml
@@ -16,7 +16,7 @@
(** Pretty-printer for Cminor *)
open Format
-open Camlcoq
+open !Camlcoq
open Integers
open AST
open PrintAST
diff --git a/backend/PrintLTL.ml b/backend/PrintLTL.ml
index d0557073..1c449e74 100644
--- a/backend/PrintLTL.ml
+++ b/backend/PrintLTL.ml
@@ -112,7 +112,7 @@ let print_function pp id f =
fprintf pp "%s() {\n" (extern_atom id);
let instrs =
List.sort
- (fun (pc1, _) (pc2, _) -> Pervasives.compare pc2 pc1)
+ (fun (pc1, _) (pc2, _) -> compare pc2 pc1)
(List.rev_map
(fun (pc, i) -> (P.to_int pc, i))
(PTree.elements f.fn_code)) in
diff --git a/backend/PrintRTL.ml b/backend/PrintRTL.ml
index ba336b0a..841540b6 100644
--- a/backend/PrintRTL.ml
+++ b/backend/PrintRTL.ml
@@ -93,7 +93,7 @@ let print_function pp id f =
fprintf pp "%s(%a) {\n" (extern_atom id) regs f.fn_params;
let instrs =
List.sort
- (fun (pc1, _) (pc2, _) -> Pervasives.compare pc2 pc1)
+ (fun (pc1, _) (pc2, _) -> compare pc2 pc1)
(List.rev_map
(fun (pc, i) -> (P.to_int pc, i))
(PTree.elements f.fn_code)) in
diff --git a/backend/PrintXTL.ml b/backend/PrintXTL.ml
index cc1f7d49..6432682a 100644
--- a/backend/PrintXTL.ml
+++ b/backend/PrintXTL.ml
@@ -138,7 +138,7 @@ let print_function pp ?alloc ?live f =
fprintf pp "f() {\n";
let instrs =
List.sort
- (fun (pc1, _) (pc2, _) -> Pervasives.compare pc2 pc1)
+ (fun (pc1, _) (pc2, _) -> compare pc2 pc1)
(List.map
(fun (pc, i) -> (P.to_int pc, i))
(PTree.elements f.fn_code)) in
diff --git a/backend/Selection.v b/backend/Selection.v
index 2d407094..4ab3331e 100644
--- a/backend/Selection.v
+++ b/backend/Selection.v
@@ -24,7 +24,7 @@
Require String.
Require Import Coqlib Maps.
-Require Import AST Errors Integers Globalenvs Switch.
+Require Import AST Errors Integers Globalenvs Builtins Switch.
Require Cminor.
Require Import Op CminorSel OpHelpers Cminortyping.
Require Import SelectOp SplitLong SelectLong SelectDiv.
@@ -162,6 +162,13 @@ Definition sel_binop (op: Cminor.binary_operation) (arg1 arg2: expr) : expr :=
| Cminor.Ocmplu c => cmplu c arg1 arg2
end.
+Definition sel_select (ty: typ) (cnd ifso ifnot: expr) : expr :=
+ let (cond, args) := condition_of_expr cnd in
+ match SelectOp.select ty cond args ifso ifnot with
+ | Some a => a
+ | None => Econdition (condexpr_of_expr cnd) ifso ifnot
+ end.
+
(** Conversion from Cminor expression to Cminorsel expressions *)
Fixpoint sel_expr (a: Cminor.expr) : expr :=
@@ -231,6 +238,43 @@ Definition sel_builtin_res (optid: option ident) : builtin_res ident :=
| Some id => BR id
end.
+(** Known builtin functions *)
+
+Function sel_known_builtin (bf: builtin_function) (args: exprlist) :=
+ match bf, args with
+ | BI_platform b, _ =>
+ SelectOp.platform_builtin b args
+ | BI_standard (BI_select ty), a1 ::: a2 ::: a3 ::: Enil =>
+ Some (sel_select ty a1 a2 a3)
+ | BI_standard BI_fabs, a1 ::: Enil =>
+ Some (SelectOp.absf a1)
+ | _, _ =>
+ None
+ end.
+
+(** Builtin functions in general *)
+
+Definition sel_builtin_default (optid: option ident) (ef: external_function)
+ (args: list Cminor.expr) :=
+ Sbuiltin (sel_builtin_res optid) ef
+ (sel_builtin_args args (Machregs.builtin_constraints ef)).
+
+Definition sel_builtin (optid: option ident) (ef: external_function)
+ (args: list Cminor.expr) :=
+ match optid, ef with
+ | Some id, EF_builtin name sg =>
+ match lookup_builtin_function name sg with
+ | Some bf =>
+ match sel_known_builtin bf (sel_exprlist args) with
+ | Some a => Sassign id a
+ | None => sel_builtin_default optid ef args
+ end
+ | None => sel_builtin_default optid ef args
+ end
+ | _, _ =>
+ sel_builtin_default optid ef args
+ end.
+
(** Conversion of Cminor [switch] statements to decision trees. *)
Parameter compile_switch: Z -> nat -> table -> comptree.
@@ -342,16 +386,10 @@ Fixpoint sel_stmt (ki: known_idents) (env: typenv) (s: Cminor.stmt) : res stmt :
OK (match classify_call fn with
| Call_default => Scall optid sg (inl _ (sel_expr fn)) (sel_exprlist args)
| Call_imm id => Scall optid sg (inr _ id) (sel_exprlist args)
- | Call_builtin ef => (Sbuiltin (sel_builtin_res optid) ef
- (sel_builtin_args args
- (Machregs.builtin_constraints ef)))
- (* sel_builtin_default optid ef args *)
- (* THIS IS WHERE TO ACTIVATE OUR OWN BUILTINS
- change sel_builtin_default to sel_builtin *)
+ | Call_builtin ef => sel_builtin optid ef args
end)
| Cminor.Sbuiltin optid ef args =>
- OK (Sbuiltin (sel_builtin_res optid) ef
- (sel_builtin_args args (Machregs.builtin_constraints ef)))
+ OK (sel_builtin optid ef args)
| Cminor.Stailcall sg fn args =>
OK (match classify_call fn with
| Call_imm id => Stailcall sg (inr _ id) (sel_exprlist args)
diff --git a/backend/Selectionaux.ml b/backend/Selectionaux.ml
index 4e366564..574c31f0 100644
--- a/backend/Selectionaux.ml
+++ b/backend/Selectionaux.ml
@@ -80,10 +80,10 @@ let fast_cmove ty =
| a, m -> failwith (Printf.sprintf "fast_cmove: unknown arch %s %s" a m)
(* The if-conversion heuristic depend on the
- -fif-conversion and -ffavor-branchless flags.
+ -fif-conversion and -Obranchless flags.
With [-fno-if-conversion] or [-0O], if-conversion is turned off entirely.
-With [-ffavor-branchless], if-conversion is performed whenever semantically
+With [-Obranchless], if-conversion is performed whenever semantically
correct, regardless of how much it could cost.
Otherwise (and by default), optimization is performed when it seems beneficial.
@@ -96,22 +96,17 @@ If-conversion seems beneficial if:
Intuition: on a modern processor, the "then" and the "else" branches
can generally be computed in parallel, there is enough ILP for that.
So, the bad case is if the most taken branch is much cheaper than the
-<<<<<<< HEAD
-other branch. Since our cost estimates are very imprecise, the
-bound on the total cost acts as a safety guard,
-=======
other branch. Another bad case is if both branches are big: since the
code for one branch precedes entirely the code for the other branch,
if the first branch contains a lot of instructions,
dynamic reordering of instructions will not look ahead far enough
to execute instructions from the other branch in parallel with
instructions from the first branch.
->>>>>>> ddb2c968e6c57d2117434f169471d87f643d831a
*)
let if_conversion_heuristic cond ifso ifnot ty =
if not !Clflags.option_fifconversion then false else
- if !Clflags.option_ffavor_branchless then true else
+ if !Clflags.option_Obranchless then true else
if not (fast_cmove ty) then false else
let c1 = cost_expr ifso and c2 = cost_expr ifnot in
c1 + c2 <= 24 && abs (c1 - c2) <= 8
diff --git a/backend/Selectionproof.v b/backend/Selectionproof.v
index 40db5d4b..8a827af2 100644
--- a/backend/Selectionproof.v
+++ b/backend/Selectionproof.v
@@ -14,7 +14,8 @@
Require Import FunInd.
Require Import Coqlib Maps.
-Require Import AST Linking Errors Integers Values Memory Events Globalenvs Smallstep.
+Require Import AST Linking Errors Integers.
+Require Import Values Memory Builtins Events Globalenvs Smallstep.
Require Import Switch Cminor Op CminorSel Cminortyping.
Require Import OpHelpers OpHelpersproof.
Require Import SelectOp SelectDiv SplitLong SelectLong Selection.
@@ -217,19 +218,16 @@ Lemma eval_condition_of_expr:
forall a le v b,
eval_expr tge sp e m le a v ->
Val.bool_of_val v b ->
- let (cond, al) := condition_of_expr a in
exists vl,
- eval_exprlist tge sp e m le al vl
- /\ eval_condition cond vl m = Some b.
+ eval_exprlist tge sp e m le (snd (condition_of_expr a)) vl
+ /\ eval_condition (fst (condition_of_expr a)) vl m = Some b.
Proof.
- intros until a. functional induction (condition_of_expr a); intros.
-(* compare *)
- inv H. simpl in H6. inv H6. apply Val.bool_of_val_of_optbool in H0.
- exists vl; auto.
-(* default *)
- exists (v :: nil); split.
- econstructor. auto. constructor.
- simpl. inv H0. auto.
+ intros a; functional induction (condition_of_expr a); intros; simpl.
+- inv H. exists vl; split; auto.
+ simpl in H6. inv H6. apply Val.bool_of_val_of_optbool in H0. auto.
+- exists (v :: nil); split.
+ constructor; auto; constructor.
+ inv H0; simpl; auto.
Qed.
Lemma eval_load:
@@ -354,6 +352,52 @@ Proof.
exists v; split; auto. eapply eval_cmplu; eauto.
Qed.
+Lemma eval_sel_select:
+ forall le a1 a2 a3 v1 v2 v3 b ty,
+ eval_expr tge sp e m le a1 v1 ->
+ eval_expr tge sp e m le a2 v2 ->
+ eval_expr tge sp e m le a3 v3 ->
+ Val.bool_of_val v1 b ->
+ exists v, eval_expr tge sp e m le (sel_select ty a1 a2 a3) v
+ /\ Val.lessdef (Val.select (Some b) v2 v3 ty) v.
+Proof.
+ unfold sel_select; intros.
+ specialize (eval_condition_of_expr _ _ _ _ H H2).
+ destruct (condition_of_expr a1) as [cond args]; simpl fst; simpl snd. intros (vl & A & B).
+ destruct (select ty cond args a2 a3) as [a|] eqn:SEL.
+- eapply eval_select; eauto.
+- exists (if b then v2 else v3); split.
+ econstructor; eauto. eapply eval_condexpr_of_expr; eauto. destruct b; auto.
+ apply Val.lessdef_normalize.
+Qed.
+
+(** Known built-in functions *)
+
+Lemma eval_sel_known_builtin:
+ forall bf args a vl v le,
+ sel_known_builtin bf args = Some a ->
+ eval_exprlist tge sp e m le args vl ->
+ builtin_function_sem bf vl = Some v ->
+ exists v', eval_expr tge sp e m le a v' /\ Val.lessdef v v'.
+Proof.
+ intros until le; intros SEL ARGS SEM.
+ destruct bf as [bf|bf]; simpl in SEL.
+- destruct bf; try discriminate.
++ (* select *)
+ inv ARGS; try discriminate. inv H0; try discriminate. inv H2; try discriminate. inv H3; try discriminate.
+ inv SEL.
+ simpl in SEM. destruct v1; inv SEM.
+ replace (Val.normalize (if Int.eq i Int.zero then v2 else v0) t)
+ with (Val.select (Some (negb (Int.eq i Int.zero))) v0 v2 t)
+ by (destruct (Int.eq i Int.zero); reflexivity).
+ eapply eval_sel_select; eauto. constructor.
++ (* fabs *)
+ inv ARGS; try discriminate. inv H0; try discriminate.
+ inv SEL.
+ simpl in SEM; inv SEM. apply eval_absf; auto.
+- eapply eval_platform_builtin; eauto.
+Qed.
+
End CMCONSTR.
(** Recognition of calls to built-in functions *)
@@ -794,6 +838,51 @@ Proof.
intros. destruct oid; simpl; auto. apply set_var_lessdef; auto.
Qed.
+Lemma sel_builtin_default_correct:
+ forall optid ef al sp e1 m1 vl t v m2 e1' m1' f k,
+ Cminor.eval_exprlist ge sp e1 m1 al vl ->
+ external_call ef ge vl m1 t v m2 ->
+ env_lessdef e1 e1' -> Mem.extends m1 m1' ->
+ exists e2' m2',
+ step tge (State f (sel_builtin_default optid ef al) k sp e1' m1')
+ t (State f Sskip k sp e2' m2')
+ /\ env_lessdef (set_optvar optid v e1) e2'
+ /\ Mem.extends m2 m2'.
+Proof.
+ intros. unfold sel_builtin_default.
+ exploit sel_builtin_args_correct; eauto. intros (vl' & A & B).
+ exploit external_call_mem_extends; eauto. intros (v' & m2' & D & E & F & _).
+ econstructor; exists m2'; split.
+ econstructor. eexact A. eapply external_call_symbols_preserved. eexact senv_preserved. eexact D.
+ split; auto. apply sel_builtin_res_correct; auto.
+Qed.
+
+Lemma sel_builtin_correct:
+ forall optid ef al sp e1 m1 vl t v m2 e1' m1' f k,
+ Cminor.eval_exprlist ge sp e1 m1 al vl ->
+ external_call ef ge vl m1 t v m2 ->
+ env_lessdef e1 e1' -> Mem.extends m1 m1' ->
+ exists e2' m2',
+ step tge (State f (sel_builtin optid ef al) k sp e1' m1')
+ t (State f Sskip k sp e2' m2')
+ /\ env_lessdef (set_optvar optid v e1) e2'
+ /\ Mem.extends m2 m2'.
+Proof.
+ intros.
+ exploit sel_exprlist_correct; eauto. intros (vl' & A & B).
+ exploit external_call_mem_extends; eauto. intros (v' & m2' & D & E & F & _).
+ unfold sel_builtin.
+ destruct optid as [id|]; eauto using sel_builtin_default_correct.
+ destruct ef; eauto using sel_builtin_default_correct.
+ destruct (lookup_builtin_function name sg) as [bf|] eqn:LKUP; eauto using sel_builtin_default_correct.
+ destruct (sel_known_builtin bf (sel_exprlist al)) as [a|] eqn:SKB; eauto using sel_builtin_default_correct.
+ simpl in D. red in D. rewrite LKUP in D. inv D.
+ exploit eval_sel_known_builtin; eauto. intros (v'' & U & V).
+ econstructor; exists m2'; split.
+ econstructor. eexact U.
+ split; auto. apply set_var_lessdef; auto. apply Val.lessdef_trans with v'; auto.
+Qed.
+
(** If-conversion *)
Lemma classify_stmt_sound_1:
@@ -983,19 +1072,18 @@ Inductive match_states: Cminor.state -> CminorSel.state -> Prop :=
match_states
(Cminor.Returnstate v k m)
(Returnstate v' k' m')
- | match_builtin_1: forall cunit hf ef args args' optid f sp e k m al f' e' k' m' env
+ | match_builtin_1: forall cunit hf ef args optid f sp e k m al f' e' k' m' env
(LINK: linkorder cunit prog)
(HF: helper_functions_declared cunit hf)
(TF: sel_function (prog_defmap cunit) hf f = OK f')
(TYF: type_function f = OK env)
(MC: match_cont cunit hf (known_id f) env k k')
- (LDA: Val.lessdef_list args args')
+ (EA: Cminor.eval_exprlist ge sp e m al args)
(LDE: env_lessdef e e')
- (ME: Mem.extends m m')
- (EA: list_forall2 (CminorSel.eval_builtin_arg tge sp e' m') al args'),
+ (ME: Mem.extends m m'),
match_states
(Cminor.Callstate (External ef) args (Cminor.Kcall optid f sp e k) m)
- (State f' (Sbuiltin (sel_builtin_res optid) ef al) k' sp e' m')
+ (State f' (sel_builtin optid ef al) k' sp e' m')
| match_builtin_2: forall cunit hf v v' optid f sp e k m f' e' m' k' env
(LINK: linkorder cunit prog)
(HF: helper_functions_declared cunit hf)
@@ -1003,11 +1091,11 @@ Inductive match_states: Cminor.state -> CminorSel.state -> Prop :=
(TYF: type_function f = OK env)
(MC: match_cont cunit hf (known_id f) env k k')
(LDV: Val.lessdef v v')
- (LDE: env_lessdef e e')
+ (LDE: env_lessdef (set_optvar optid v e) e')
(ME: Mem.extends m m'),
match_states
(Cminor.Returnstate v (Cminor.Kcall optid f sp e k) m)
- (State f' Sskip k' sp (set_builtin_res (sel_builtin_res optid) v' e') m').
+ (State f' Sskip k' sp e' m').
Remark call_cont_commut:
forall cunit hf ki env k k',
@@ -1079,6 +1167,14 @@ Proof.
destruct (ident_eq id id0). conclude. discriminate.
Qed.
+Remark sel_builtin_nolabel:
+ forall (hf: helper_functions) optid ef args, nolabel' (sel_builtin optid ef args).
+Proof.
+ unfold sel_builtin; intros; red; intros.
+ destruct optid; auto. destruct ef; auto. destruct lookup_builtin_function; auto.
+ destruct sel_known_builtin; auto.
+Qed.
+
Remark find_label_commut:
forall cunit hf ki env lbl s k s' k',
match_cont cunit hf ki env k k' ->
@@ -1094,8 +1190,11 @@ Proof.
- unfold store. destruct (addressing m (sel_expr e)); simpl; auto.
(* call *)
- destruct (classify_call (prog_defmap cunit) e); simpl; auto.
+ rewrite sel_builtin_nolabel; auto.
(* tailcall *)
- destruct (classify_call (prog_defmap cunit) e); simpl; auto.
+(* builtin *)
+- rewrite sel_builtin_nolabel; auto.
(* seq *)
- exploit (IHs1 (Cminor.Kseq s2 k)). constructor; eauto. eauto.
destruct (Cminor.find_label lbl s1 (Cminor.Kseq s2 k)) as [[sx kx] | ];
@@ -1189,9 +1288,7 @@ Proof.
eapply match_cont_call with (cunit := cunit) (hf := hf); eauto.
+ (* turned into Sbuiltin *)
intros EQ. subst fd.
- exploit sel_builtin_args_correct; eauto. intros [vargs' [C D]].
- right; left; split. simpl. omega. split. auto.
- econstructor; eauto.
+ right; left; split. simpl; omega. split; auto. econstructor; eauto.
- (* Stailcall *)
exploit Mem.free_parallel_extends; eauto. intros [m2' [P Q]].
erewrite <- stackspace_function_translated in P by eauto.
@@ -1208,12 +1305,8 @@ Proof.
eapply match_callstate with (cunit := cunit'); eauto.
eapply call_cont_commut; eauto.
- (* Sbuiltin *)
- exploit sel_builtin_args_correct; eauto. intros [vargs' [P Q]].
- exploit external_call_mem_extends; eauto.
- intros [vres' [m2 [A [B [C D]]]]].
- left; econstructor; split.
- econstructor. eauto. eapply external_call_symbols_preserved; eauto. apply senv_preserved.
- econstructor; eauto. apply sel_builtin_res_correct; auto.
+ exploit sel_builtin_correct; eauto. intros (e2' & m2' & P & Q & R).
+ left; econstructor; split. eexact P. econstructor; eauto.
- (* Seq *)
left; econstructor; split.
constructor.
@@ -1304,11 +1397,8 @@ Proof.
econstructor. eapply external_call_symbols_preserved; eauto. apply senv_preserved.
econstructor; eauto.
- (* external call turned into a Sbuiltin *)
- exploit external_call_mem_extends; eauto.
- intros [vres' [m2 [A [B [C D]]]]].
- left; econstructor; split.
- econstructor. eauto. eapply external_call_symbols_preserved; eauto. apply senv_preserved.
- econstructor; eauto.
+ exploit sel_builtin_correct; eauto. intros (e2' & m2' & P & Q & R).
+ left; econstructor; split. eexact P. econstructor; eauto.
- (* return *)
inv MC.
left; econstructor; split.
@@ -1316,7 +1406,6 @@ Proof.
econstructor; eauto. destruct optid; simpl; auto. apply set_var_lessdef; auto.
- (* return of an external call turned into a Sbuiltin *)
right; left; split. simpl; omega. split. auto. econstructor; eauto.
- apply sel_builtin_res_correct; auto.
Qed.
Lemma sel_initial_states:
diff --git a/backend/SplitLongproof.v b/backend/SplitLongproof.v
index 6718ba5b..c8e3b94c 100644
--- a/backend/SplitLongproof.v
+++ b/backend/SplitLongproof.v
@@ -15,13 +15,16 @@
Require Import String.
Require Import Coqlib Maps.
Require Import AST Errors Integers Floats.
-Require Import Values Memory Globalenvs Events Cminor Op CminorSel.
+Require Import Values Memory Globalenvs Builtins Events Cminor Op CminorSel.
Require Import OpHelpers OpHelpersproof.
+Require Import Values Memory Globalenvs Builtins Events Cminor Op CminorSel.
Require Import SelectOp SelectOpproof SplitLong.
Local Open Scope cminorsel_scope.
Local Open Scope string_scope.
+(** * Properties of the helper functions *)
+
(** * Correctness of the instruction selection functions for 64-bit operators *)
Section CMCONSTR.
@@ -38,56 +41,64 @@ Ltac UseHelper := decompose [Logic.and] arith_helpers_correct; eauto.
Ltac DeclHelper := red in HELPERS; decompose [Logic.and] HELPERS; eauto.
Lemma eval_helper:
- forall le id name sg args vargs vres,
+ forall bf le id name sg args vargs vres,
eval_exprlist ge sp e m le args vargs ->
helper_declared prog id name sg ->
- external_implements name sg vargs vres ->
+ lookup_builtin_function name sg = Some bf ->
+ builtin_function_sem bf vargs = Some vres ->
eval_expr ge sp e m le (Eexternal id sg args) vres.
Proof.
intros.
red in H0. apply Genv.find_def_symbol in H0. destruct H0 as (b & P & Q).
rewrite <- Genv.find_funct_ptr_iff in Q.
- econstructor; eauto.
+ econstructor; eauto.
+ simpl. red. rewrite H1. constructor; auto.
Qed.
Corollary eval_helper_1:
- forall le id name sg arg1 varg1 vres,
+ forall bf le id name sg arg1 varg1 vres,
eval_expr ge sp e m le arg1 varg1 ->
helper_declared prog id name sg ->
- external_implements name sg (varg1::nil) vres ->
+ lookup_builtin_function name sg = Some bf ->
+ builtin_function_sem bf (varg1 :: nil) = Some vres ->
eval_expr ge sp e m le (Eexternal id sg (arg1 ::: Enil)) vres.
Proof.
intros. eapply eval_helper; eauto. constructor; auto. constructor.
Qed.
Corollary eval_helper_2:
- forall le id name sg arg1 arg2 varg1 varg2 vres,
+ forall bf le id name sg arg1 arg2 varg1 varg2 vres,
eval_expr ge sp e m le arg1 varg1 ->
eval_expr ge sp e m le arg2 varg2 ->
helper_declared prog id name sg ->
- external_implements name sg (varg1::varg2::nil) vres ->
+ lookup_builtin_function name sg = Some bf ->
+ builtin_function_sem bf (varg1 :: varg2 :: nil) = Some vres ->
eval_expr ge sp e m le (Eexternal id sg (arg1 ::: arg2 ::: Enil)) vres.
Proof.
intros. eapply eval_helper; eauto. constructor; auto. constructor; auto. constructor.
Qed.
Remark eval_builtin_1:
- forall le id sg arg1 varg1 vres,
+ forall bf le id sg arg1 varg1 vres,
eval_expr ge sp e m le arg1 varg1 ->
- builtin_implements id sg (varg1::nil) vres ->
+ lookup_builtin_function id sg = Some bf ->
+ builtin_function_sem bf (varg1 :: nil) = Some vres ->
eval_expr ge sp e m le (Ebuiltin (EF_builtin id sg) (arg1 ::: Enil)) vres.
Proof.
- intros. econstructor. econstructor. eauto. constructor. apply H0.
+ intros. econstructor. econstructor. eauto. constructor.
+ simpl. red. rewrite H0. constructor. auto.
Qed.
Remark eval_builtin_2:
- forall le id sg arg1 arg2 varg1 varg2 vres,
+ forall bf le id sg arg1 arg2 varg1 varg2 vres,
eval_expr ge sp e m le arg1 varg1 ->
eval_expr ge sp e m le arg2 varg2 ->
- builtin_implements id sg (varg1::varg2::nil) vres ->
+ lookup_builtin_function id sg = Some bf ->
+ builtin_function_sem bf (varg1 :: varg2 :: nil) = Some vres ->
eval_expr ge sp e m le (Ebuiltin (EF_builtin id sg) (arg1 ::: arg2 ::: Enil)) vres.
Proof.
- intros. econstructor. constructor; eauto. constructor; eauto. constructor. apply H1.
+ intros. econstructor. constructor; eauto. constructor; eauto. constructor.
+ simpl. red. rewrite H1. constructor. auto.
Qed.
Definition unary_constructor_sound (cstr: expr -> expr) (sem: val -> val) : Prop :=
@@ -336,9 +347,10 @@ Qed.
Theorem eval_negl: unary_constructor_sound negl Val.negl.
Proof.
unfold negl; red; intros. destruct (is_longconst a) eqn:E.
- econstructor; split. apply eval_longconst.
+- econstructor; split. apply eval_longconst.
exploit is_longconst_sound; eauto. intros EQ; subst x. simpl. auto.
- econstructor; split. eapply eval_builtin_1; eauto. UseHelper. auto.
+- exists (Val.negl x); split; auto.
+ eapply (eval_builtin_1 (BI_standard BI_negl)); eauto.
Qed.
Theorem eval_notl: unary_constructor_sound notl Val.notl.
@@ -360,7 +372,7 @@ Theorem eval_longoffloat:
exists v, eval_expr ge sp e m le (longoffloat a) v /\ Val.lessdef y v.
Proof.
intros; unfold longoffloat. econstructor; split.
- eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
+ eapply (eval_helper_1 (BI_standard BI_i64_dtos)); eauto. DeclHelper. auto. auto.
Qed.
Theorem eval_longuoffloat:
@@ -370,7 +382,7 @@ Theorem eval_longuoffloat:
exists v, eval_expr ge sp e m le (longuoffloat a) v /\ Val.lessdef y v.
Proof.
intros; unfold longuoffloat. econstructor; split.
- eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
+ eapply (eval_helper_1 (BI_standard BI_i64_dtou)); eauto. DeclHelper. auto. auto.
Qed.
Theorem eval_floatoflong:
@@ -379,8 +391,9 @@ Theorem eval_floatoflong:
Val.floatoflong x = Some y ->
exists v, eval_expr ge sp e m le (floatoflong a) v /\ Val.lessdef y v.
Proof.
- intros; unfold floatoflong. econstructor; split.
- eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
+ intros; unfold floatoflong. exists y; split; auto.
+ eapply (eval_helper_1 (BI_standard BI_i64_stod)); eauto. DeclHelper. auto.
+ simpl. destruct x; simpl in H0; inv H0; auto.
Qed.
Theorem eval_floatoflongu:
@@ -389,8 +402,9 @@ Theorem eval_floatoflongu:
Val.floatoflongu x = Some y ->
exists v, eval_expr ge sp e m le (floatoflongu a) v /\ Val.lessdef y v.
Proof.
- intros; unfold floatoflongu. econstructor; split.
- eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
+ intros; unfold floatoflongu. exists y; split; auto.
+ eapply (eval_helper_1 (BI_standard BI_i64_utod)); eauto. DeclHelper. auto.
+ simpl. destruct x; simpl in H0; inv H0; auto.
Qed.
Theorem eval_longofsingle:
@@ -427,8 +441,9 @@ Theorem eval_singleoflong:
Val.singleoflong x = Some y ->
exists v, eval_expr ge sp e m le (singleoflong a) v /\ Val.lessdef y v.
Proof.
- intros; unfold singleoflong. econstructor; split.
- eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
+ intros; unfold singleoflong. exists y; split; auto.
+ eapply (eval_helper_1 (BI_standard BI_i64_stof)); eauto. DeclHelper. auto.
+ simpl. destruct x; simpl in H0; inv H0; auto.
Qed.
Theorem eval_singleoflongu:
@@ -437,8 +452,9 @@ Theorem eval_singleoflongu:
Val.singleoflongu x = Some y ->
exists v, eval_expr ge sp e m le (singleoflongu a) v /\ Val.lessdef y v.
Proof.
- intros; unfold singleoflongu. econstructor; split.
- eapply eval_helper_1; eauto. DeclHelper. UseHelper. auto.
+ intros; unfold singleoflongu. exists y; split; auto.
+ eapply (eval_helper_1 (BI_standard BI_i64_utof)); eauto. DeclHelper. auto.
+ simpl. destruct x; simpl in H0; inv H0; auto.
Qed.
Theorem eval_andl: binary_constructor_sound andl Val.andl.
@@ -565,7 +581,9 @@ Proof.
simpl. erewrite <- Int64.decompose_shl_2. instantiate (1 := Int64.hiword i).
rewrite Int64.ofwords_recompose. auto. auto.
+ (* n >= 64 *)
- econstructor; split. eapply eval_helper_2; eauto. EvalOp. DeclHelper. UseHelper. auto.
+ econstructor; split.
+ eapply eval_helper_2; eauto. EvalOp. DeclHelper. reflexivity. reflexivity.
+ auto.
Qed.
Theorem eval_shll: binary_constructor_sound shll Val.shll.
@@ -576,7 +594,7 @@ Proof.
exploit is_intconst_sound; eauto. intros EQ; subst y; clear H0.
eapply eval_shllimm; eauto.
- (* General case *)
- econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto.
+ econstructor; split. eapply eval_helper_2; eauto. DeclHelper. reflexivity. reflexivity. auto.
Qed.
Lemma eval_shrluimm:
@@ -610,7 +628,9 @@ Proof.
simpl. erewrite <- Int64.decompose_shru_2. instantiate (1 := Int64.loword i).
rewrite Int64.ofwords_recompose. auto. auto.
+ (* n >= 64 *)
- econstructor; split. eapply eval_helper_2; eauto. EvalOp. DeclHelper. UseHelper. auto.
+ econstructor; split.
+ eapply eval_helper_2; eauto. EvalOp. DeclHelper. reflexivity. reflexivity.
+ auto.
Qed.
Theorem eval_shrlu: binary_constructor_sound shrlu Val.shrlu.
@@ -621,7 +641,7 @@ Proof.
exploit is_intconst_sound; eauto. intros EQ; subst y; clear H0.
eapply eval_shrluimm; eauto.
- (* General case *)
- econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto.
+ econstructor; split. eapply eval_helper_2; eauto. DeclHelper. reflexivity. reflexivity. auto.
Qed.
Lemma eval_shrlimm:
@@ -659,7 +679,9 @@ Proof.
erewrite <- Int64.decompose_shr_2. instantiate (1 := Int64.loword i).
rewrite Int64.ofwords_recompose. auto. auto.
+ (* n >= 64 *)
- econstructor; split. eapply eval_helper_2; eauto. EvalOp. DeclHelper. UseHelper. auto.
+ econstructor; split.
+ eapply eval_helper_2; eauto. EvalOp. DeclHelper. reflexivity. reflexivity.
+ auto.
Qed.
Theorem eval_shrl: binary_constructor_sound shrl Val.shrl.
@@ -670,7 +692,7 @@ Proof.
exploit is_intconst_sound; eauto. intros EQ; subst y; clear H0.
eapply eval_shrlimm; eauto.
- (* General case *)
- econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto.
+ econstructor; split. eapply eval_helper_2; eauto. DeclHelper. reflexivity. reflexivity. auto.
Qed.
Theorem eval_addl: Archi.ptr64 = false -> binary_constructor_sound addl Val.addl.
@@ -680,7 +702,7 @@ Proof.
assert (DEFAULT:
exists v, eval_expr ge sp e m le default v /\ Val.lessdef (Val.addl x y) v).
{
- econstructor; split. eapply eval_builtin_2; eauto. UseHelper. auto.
+ econstructor; split. eapply eval_builtin_2; eauto. reflexivity. reflexivity. auto.
}
destruct (is_longconst a) as [p|] eqn:LC1;
destruct (is_longconst b) as [q|] eqn:LC2.
@@ -703,7 +725,7 @@ Proof.
assert (DEFAULT:
exists v, eval_expr ge sp e m le default v /\ Val.lessdef (Val.subl x y) v).
{
- econstructor; split. eapply eval_builtin_2; eauto. UseHelper. auto.
+ econstructor; split. eapply eval_builtin_2; eauto. reflexivity. reflexivity. auto.
}
destruct (is_longconst a) as [p|] eqn:LC1;
destruct (is_longconst b) as [q|] eqn:LC2.
@@ -734,7 +756,7 @@ Proof.
exploit eval_add. eexact E2. eexact E3. intros [v5 [E5 L5]].
exploit eval_add. eexact E5. eexact E4. intros [v6 [E6 L6]].
exists (Val.longofwords v6 (Val.loword p)); split.
- EvalOp. eapply eval_builtin_2; eauto. UseHelper.
+ EvalOp. eapply eval_builtin_2; eauto. reflexivity. reflexivity.
intros. unfold le1, p in *; subst; simpl in *.
inv L3. inv L4. inv L5. simpl in L6. inv L6.
simpl. f_equal. symmetry. apply Int64.decompose_mul.
@@ -782,14 +804,14 @@ Theorem eval_mullhu:
forall n, unary_constructor_sound (fun a => mullhu a n) (fun v => Val.mullhu v (Vlong n)).
Proof.
unfold mullhu; intros; red; intros. econstructor; split; eauto.
- eapply eval_helper_2; eauto. apply eval_longconst. DeclHelper; eauto. UseHelper.
+ eapply eval_helper_2; eauto. apply eval_longconst. DeclHelper. reflexivity. reflexivity.
Qed.
Theorem eval_mullhs:
forall n, unary_constructor_sound (fun a => mullhs a n) (fun v => Val.mullhs v (Vlong n)).
Proof.
unfold mullhs; intros; red; intros. econstructor; split; eauto.
- eapply eval_helper_2; eauto. apply eval_longconst. DeclHelper; eauto. UseHelper.
+ eapply eval_helper_2; eauto. apply eval_longconst. DeclHelper. reflexivity. reflexivity.
Qed.
Theorem eval_shrxlimm:
@@ -831,7 +853,7 @@ Theorem eval_divlu_base:
exists v, eval_expr ge sp e m le (divlu_base a b) v /\ Val.lessdef z v.
Proof.
intros; unfold divlu_base.
- econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto.
+ econstructor; split. eapply eval_helper_2; eauto. DeclHelper. reflexivity. eassumption. auto.
Qed.
Theorem eval_modlu_base:
@@ -842,7 +864,7 @@ Theorem eval_modlu_base:
exists v, eval_expr ge sp e m le (modlu_base a b) v /\ Val.lessdef z v.
Proof.
intros; unfold modlu_base.
- econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto.
+ econstructor; split. eapply eval_helper_2; eauto. DeclHelper. reflexivity. eassumption. auto.
Qed.
Theorem eval_divls_base:
@@ -853,7 +875,7 @@ Theorem eval_divls_base:
exists v, eval_expr ge sp e m le (divls_base a b) v /\ Val.lessdef z v.
Proof.
intros; unfold divls_base.
- econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto.
+ econstructor; split. eapply eval_helper_2; eauto. DeclHelper. reflexivity. eassumption. auto.
Qed.
Theorem eval_modls_base:
@@ -864,7 +886,7 @@ Theorem eval_modls_base:
exists v, eval_expr ge sp e m le (modls_base a b) v /\ Val.lessdef z v.
Proof.
intros; unfold modls_base.
- econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto.
+ econstructor; split. eapply eval_helper_2; eauto. DeclHelper. reflexivity. eassumption. auto.
Qed.
Remark decompose_cmpl_eq_zero:
diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v
index 605d7e90..5072448a 100644
--- a/backend/ValueAnalysis.v
+++ b/backend/ValueAnalysis.v
@@ -13,7 +13,7 @@
Require Import FunInd.
Require Import Coqlib Maps Integers Floats Lattice Kildall.
Require Import Compopts AST Linking.
-Require Import Values Memory Globalenvs Events.
+Require Import Values Memory Globalenvs Builtins Events.
Require Import Registers Op RTL.
Require Import ValueDomain ValueAOp Liveness.
@@ -78,6 +78,15 @@ Definition transfer_builtin_default
let (av, am') := analyze_call am (map (abuiltin_arg ae am rm) args) in
VA.State (set_builtin_res res av ae) am'.
+Definition eval_static_builtin_function
+ (ae: aenv) (am: amem) (rm: romem)
+ (bf: builtin_function) (args: list (builtin_arg reg)) :=
+ match builtin_function_sem bf
+ (map val_of_aval (map (abuiltin_arg ae am rm) args)) with
+ | Some v => aval_of_val v
+ | None => None
+ end.
+
Definition transfer_builtin
(ae: aenv) (am: amem) (rm: romem) (ef: external_function)
(args: list (builtin_arg reg)) (res: builtin_res reg) :=
@@ -105,6 +114,15 @@ Definition transfer_builtin
| EF_annot_val _ _ _, v :: nil =>
let av := abuiltin_arg ae am rm v in
VA.State (set_builtin_res res av ae) am
+ | EF_builtin name sg, _ =>
+ match lookup_builtin_function name sg with
+ | Some bf =>
+ match eval_static_builtin_function ae am rm bf args with
+ | Some av => VA.State (set_builtin_res res av ae) am
+ | None => transfer_builtin_default ae am rm args res
+ end
+ | None => transfer_builtin_default ae am rm args res
+ end
| _, _ =>
transfer_builtin_default ae am rm args res
end.
@@ -372,6 +390,31 @@ Proof.
intros. destruct res; simpl; auto. apply ematch_update; auto.
Qed.
+Lemma eval_static_builtin_function_sound:
+ forall bc ge rs sp m ae rm am (bf: builtin_function) al vl v va,
+ ematch bc rs ae ->
+ romatch bc m rm ->
+ mmatch bc m am ->
+ genv_match bc ge ->
+ bc sp = BCstack ->
+ eval_builtin_args ge (fun r => rs#r) (Vptr sp Ptrofs.zero) m al vl ->
+ eval_static_builtin_function ae am rm bf al = Some va ->
+ builtin_function_sem bf vl = Some v ->
+ vmatch bc v va.
+Proof.
+ unfold eval_static_builtin_function; intros.
+ exploit abuiltin_args_sound; eauto.
+ set (vla := map (abuiltin_arg ae am rm) al) in *. intros VMA.
+ destruct (builtin_function_sem bf (map val_of_aval vla)) as [v0|] eqn:A; try discriminate.
+ assert (LD: Val.lessdef v0 v).
+ { apply val_inject_lessdef.
+ exploit (bs_inject _ (builtin_function_sem bf)).
+ apply val_inject_list_lessdef. eapply list_val_of_aval_sound; eauto.
+ rewrite A, H6; simpl. auto.
+ }
+ inv LD. apply aval_of_val_sound; auto. discriminate.
+Qed.
+
(** ** Constructing block classifications *)
Definition bc_nostack (bc: block_classification) : Prop :=
@@ -1343,6 +1386,13 @@ Proof.
}
unfold transfer_builtin in TR.
destruct ef; auto.
++ (* builtin function *)
+ destruct (lookup_builtin_function name sg) as [bf|] eqn:LK; auto.
+ destruct (eval_static_builtin_function ae am rm bf args) as [av|] eqn:ES; auto.
+ simpl in H1. red in H1. rewrite LK in H1. inv H1.
+ eapply sound_succ_state; eauto. simpl; auto.
+ apply set_builtin_res_sound; auto.
+ eapply eval_static_builtin_function_sound; eauto.
+ (* volatile load *)
inv H0; auto. inv H3; auto. inv H1.
exploit abuiltin_arg_sound; eauto. intros VM1.
diff --git a/backend/ValueDomain.v b/backend/ValueDomain.v
index f6afa836..fd3bd5ae 100644
--- a/backend/ValueDomain.v
+++ b/backend/ValueDomain.v
@@ -13,7 +13,7 @@
Require Import FunInd.
Require Import Zwf Coqlib Maps Zbits Integers Floats Lattice.
Require Import Compopts AST.
-Require Import Values Memory Globalenvs Events.
+Require Import Values Memory Globalenvs Builtins Events.
Require Import Registers RTL.
(** The abstract domains for value analysis *)
@@ -3038,7 +3038,47 @@ Proof with (auto using provenance_monotone with va).
- destruct (zlt n 16)...
Qed.
-(** Abstracting memory blocks *)
+(** Analysis of known builtin functions. All we have is a dynamic semantics
+ as a function [list val -> option val], but we can still perform
+ some constant propagation. *)
+
+Definition val_of_aval (a: aval) : val :=
+ match a with
+ | I n => Vint n
+ | L n => Vlong n
+ | F f => Vfloat f
+ | FS f => Vsingle f
+ | _ => Vundef
+ end.
+
+Definition aval_of_val (v: val) : option aval :=
+ match v with
+ | Vint n => Some (I n)
+ | Vlong n => Some (L n)
+ | Vfloat f => Some (F f)
+ | Vsingle f => Some (FS f)
+ | _ => None
+ end.
+
+Lemma val_of_aval_sound:
+ forall v a, vmatch v a -> Val.lessdef (val_of_aval a) v.
+Proof.
+ destruct 1; simpl; auto.
+Qed.
+
+Corollary list_val_of_aval_sound:
+ forall vl al, list_forall2 vmatch vl al -> Val.lessdef_list (map val_of_aval al) vl.
+Proof.
+ induction 1; simpl; constructor; auto using val_of_aval_sound.
+Qed.
+
+Lemma aval_of_val_sound:
+ forall v a, aval_of_val v = Some a -> vmatch v a.
+Proof.
+ intros v a E; destruct v; simpl in E; inv E; constructor.
+Qed.
+
+(** * Abstracting memory blocks *)
Inductive acontent : Type :=
| ACval (chunk: memory_chunk) (av: aval).
diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml
index 37527940..dc25b516 100644
--- a/cfrontend/C2C.ml
+++ b/cfrontend/C2C.ml
@@ -153,19 +153,17 @@ let ais_annot_functions =
true);]
else
[]
-
-let builtin_ternary suffix typ =
- ("__builtin_ternary_" ^ suffix),
- (typ, [TInt(IInt, []); typ; typ], false);;
let builtins_generic = {
- Builtins.typedefs = [];
- Builtins.functions =
+ builtin_typedefs = [];
+ builtin_functions =
ais_annot_functions
@
[
(* Integer arithmetic *)
- "__builtin_bswap",
+ "__builtin_bswap64",
+ (TInt(IULongLong, []), [TInt(IULongLong, [])], false);
+ "__builtin_bswap",
(TInt(IUInt, []), [TInt(IUInt, [])], false);
"__builtin_bswap32",
(TInt(IUInt, []), [TInt(IUInt, [])], false);
@@ -184,15 +182,12 @@ let builtins_generic = {
TPtr(TVoid [AConst], []);
TInt(IULong, []);
TInt(IULong, [])],
- false);
- (* Ternary operator *)
- builtin_ternary "uint" (TInt(IUInt, []));
- builtin_ternary "ulong" (TInt(IULong, []));
- builtin_ternary "int" (TInt(IInt, []));
- builtin_ternary "long" (TInt(ILong, []));
- builtin_ternary "double" (TFloat(FDouble, []));
- builtin_ternary "float" (TFloat(FFloat, []));
-
+ false);
+ (* Selection *)
+ "__builtin_sel",
+ (TVoid [],
+ [TInt(C.IBool, [])],
+ true);
(* Annotations *)
"__builtin_annot",
(TVoid [],
@@ -336,9 +331,12 @@ let builtins_generic = {
(* Add processor-dependent builtins *)
-let builtins =
- Builtins.({ typedefs = builtins_generic.typedefs @ CBuiltins.builtins.typedefs;
- functions = builtins_generic.Builtins.functions @ CBuiltins.builtins.functions })
+let builtins = {
+ builtin_typedefs =
+ builtins_generic.builtin_typedefs @ CBuiltins.builtins.builtin_typedefs;
+ builtin_functions =
+ builtins_generic.builtin_functions @ CBuiltins.builtins.builtin_functions
+}
(** ** The known attributes *)
@@ -632,6 +630,12 @@ and convertParams env = function
| [] -> Tnil
| (id, ty) :: rem -> Tcons(convertTyp env ty, convertParams env rem)
+(* Convert types for the arguments to a function call. The types for
+ fixed arguments are taken from the function prototype. The types
+ for other arguments (variable-argument function or unprototyped K&R
+ functions) are taken from the types of the function arguments,
+ after default argument conversion. *)
+
let rec convertTypArgs env tl el =
match tl, el with
| _, [] -> Tnil
@@ -641,6 +645,20 @@ let rec convertTypArgs env tl el =
| (id, t1) :: tl, e1 :: el ->
Tcons(convertTyp env t1, convertTypArgs env tl el)
+(* Convert types for the arguments to inline asm statements and to
+ the special built-in functions __builtin_annot, __builtin_ais_annot_
+ and __builtin_debug. The types are taken from the types of the
+ arguments, after performing the usual unary conversions.
+ Hence char becomes int but float remains float and is not promoted
+ to double. The goal is to preserve the representation of the arguments
+ and avoid inserting compiled code to convert the arguments. *)
+
+let rec convertTypAnnotArgs env = function
+ | [] -> Tnil
+ | e1 :: el ->
+ Tcons(convertTyp env (Cutil.unary_conversion env e1.etyp),
+ convertTypAnnotArgs env el)
+
let convertField env f =
if f.fld_bitfield <> None then
unsupported "bit field in struct or union (consider adding option [-fbitfields])";
@@ -881,7 +899,7 @@ let rec convertExpr env e =
| {edesc = C.EVar id} :: args2 -> (id.name, args2)
| _::args2 -> error "argument 2 of '__builtin_debug' must be either a string literal or a variable"; ("", args2)
| [] -> assert false (* catched earlier *) in
- let targs2 = convertTypArgs env [] args2 in
+ let targs2 = convertTypAnnotArgs env args2 in
Ebuiltin(
AST.EF_debug(P.of_int64 kind, intern_string text,
typlist_of_typelist targs2),
@@ -890,7 +908,7 @@ let rec convertExpr env e =
| C.ECall({edesc = C.EVar {name = "__builtin_annot"}}, args) ->
begin match args with
| {edesc = C.EConst(CStr txt)} :: args1 ->
- let targs1 = convertTypArgs env [] args1 in
+ let targs1 = convertTypAnnotArgs env args1 in
Ebuiltin(
AST.EF_annot(P.of_int 1,coqstring_of_camlstring txt, typlist_of_typelist targs1),
targs1, convertExprList env args1, convertTyp env e.etyp)
@@ -918,7 +936,7 @@ let rec convertExpr env e =
let file,line = !currentLocation in
let fun_name = !currentFunction in
let loc_string = Printf.sprintf "# file:%s line:%d function:%s\n" file line fun_name in
- let targs1 = convertTypArgs env [] args1 in
+ let targs1 = convertTypAnnotArgs env args1 in
AisAnnot.validate_ais_annot env !currentLocation txt args1;
Ebuiltin(
AST.EF_annot(P.of_int 2,coqstring_of_camlstring (loc_string ^ txt), typlist_of_typelist targs1),
@@ -954,6 +972,10 @@ let rec convertExpr env e =
Econs(va_list_ptr dst, Econs(va_list_ptr src, Enil)),
Tvoid)
+ | C.ECall({edesc = C.EVar {name = "__builtin_sel"}}, [arg1; arg2; arg3]) ->
+ ewrap (Ctyping.eselection (convertExpr env arg1)
+ (convertExpr env arg2) (convertExpr env arg3))
+
| C.ECall({edesc = C.EVar {name = "printf"}}, args)
when !Clflags.option_interp ->
let targs = convertTypArgs env [] args
@@ -1019,14 +1041,14 @@ let convertAsm loc env txt outputs inputs clobber =
match output' with None -> TVoid [] | Some e -> e.etyp in
(* Build the Ebuiltin expression *)
let e =
- let tinputs = convertTypArgs env [] inputs' in
+ let tinputs = convertTypAnnotArgs env inputs' in
let toutput = convertTyp env ty_res in
Ebuiltin( AST.EF_inline_asm(coqstring_of_camlstring txt',
signature_of_type tinputs toutput AST.cc_default,
clobber'),
tinputs,
convertExprList env inputs',
- convertTyp env ty_res) in
+ toutput) in
(* Add an assignment to the output, if any *)
match output' with
| None -> e
@@ -1257,7 +1279,7 @@ let convertFundecl env (sto, id, ty, optinit) =
then AST.EF_runtime(id'', sg)
else
if Str.string_match re_builtin id.name 0
- && List.mem_assoc id.name builtins.Builtins.functions
+ && List.mem_assoc id.name builtins.builtin_functions
then AST.EF_builtin(id'', sg)
else AST.EF_external(id'', sg) in
(id', AST.Gfun(Ctypes.External(ef, args, res, cconv)))
@@ -1456,7 +1478,7 @@ let convertProgram p =
Hashtbl.clear decl_atom;
Hashtbl.clear stringTable;
Hashtbl.clear wstringTable;
- let p = cleanupGlobals (Builtins.declarations() @ p) in
+ let p = cleanupGlobals (Env.initial_declarations() @ p) in
try
let env = translEnv Env.empty p in
let typs = convertCompositedefs env [] p in
diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v
index 7f5fe355..2942080b 100644
--- a/cfrontend/Cexec.v
+++ b/cfrontend/Cexec.v
@@ -16,7 +16,7 @@ Require Import FunInd.
Require Import Axioms Classical.
Require Import String Coqlib Decidableplus.
Require Import Errors Maps Integers Floats.
-Require Import AST Values Memory Events Globalenvs Determinism.
+Require Import AST Values Memory Events Globalenvs Builtins Determinism.
Require Import Ctypes Cop Csyntax Csem.
Require Cstrategy.
@@ -501,12 +501,19 @@ Definition do_ef_debug (kind: positive) (text: ident) (targs: list typ)
(w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
Some(w, E0, Vundef, m).
+Definition do_builtin_or_external (name: string) (sg: signature)
+ (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
+ match lookup_builtin_function name sg with
+ | Some bf => match builtin_function_sem bf vargs with Some v => Some(w, E0, v, m) | None => None end
+ | None => do_external_function name sg ge w vargs m
+ end.
+
Definition do_external (ef: external_function):
world -> list val -> mem -> option (world * trace * val * mem) :=
match ef with
| EF_external name sg => do_external_function name sg ge
- | EF_builtin name sg => do_external_function name sg ge
- | EF_runtime name sg => do_external_function name sg ge
+ | EF_builtin name sg => do_builtin_or_external name sg
+ | EF_runtime name sg => do_builtin_or_external name sg
| EF_vload chunk => do_ef_volatile_load chunk
| EF_vstore chunk => do_ef_volatile_store chunk
| EF_malloc => do_ef_malloc
@@ -523,17 +530,26 @@ Lemma do_ef_external_sound:
do_external ef w vargs m = Some(w', t, vres, m') ->
external_call ef ge vargs m t vres m' /\ possible_trace w t w'.
Proof with try congruence.
+ intros until m'.
assert (SIZE: forall v sz, do_alloc_size v = Some sz -> v = Vptrofs sz).
{ intros until sz; unfold Vptrofs; destruct v; simpl; destruct Archi.ptr64 eqn:SF;
intros EQ; inv EQ; f_equal; symmetry; eauto with ptrofs. }
- intros until m'.
+ assert (BF_EX: forall name sg,
+ do_builtin_or_external name sg w vargs m = Some (w', t, vres, m') ->
+ builtin_or_external_sem name sg ge vargs m t vres m' /\ possible_trace w t w').
+ { unfold do_builtin_or_external, builtin_or_external_sem; intros.
+ destruct (lookup_builtin_function name sg ) as [bf|].
+ - destruct (builtin_function_sem bf vargs) as [vres1|] eqn:BF; inv H.
+ split. constructor; auto. constructor.
+ - eapply do_external_function_sound; eauto.
+ }
destruct ef; simpl.
(* EF_external *)
eapply do_external_function_sound; eauto.
(* EF_builtin *)
- eapply do_external_function_sound; eauto.
+ eapply BF_EX; eauto.
(* EF_runtime *)
- eapply do_external_function_sound; eauto.
+ eapply BF_EX; eauto.
(* EF_vload *)
unfold do_ef_volatile_load. destruct vargs... destruct v... destruct vargs...
mydestr. destruct p as [[w'' t''] v]; mydestr.
@@ -575,17 +591,26 @@ Lemma do_ef_external_complete:
external_call ef ge vargs m t vres m' -> possible_trace w t w' ->
do_external ef w vargs m = Some(w', t, vres, m').
Proof.
+ intros.
assert (SIZE: forall n, do_alloc_size (Vptrofs n) = Some n).
{ unfold Vptrofs, do_alloc_size; intros; destruct Archi.ptr64 eqn:SF.
rewrite Ptrofs.of_int64_to_int64; auto.
rewrite Ptrofs.of_int_to_int; auto. }
- intros. destruct ef; simpl in *.
+ assert (BF_EX: forall name sg,
+ builtin_or_external_sem name sg ge vargs m t vres m' ->
+ do_builtin_or_external name sg w vargs m = Some (w', t, vres, m')).
+ { unfold do_builtin_or_external, builtin_or_external_sem; intros.
+ destruct (lookup_builtin_function name sg) as [bf|].
+ - inv H1. inv H0. rewrite H2. auto.
+ - eapply do_external_function_complete; eauto.
+ }
+ destruct ef; simpl in *.
(* EF_external *)
eapply do_external_function_complete; eauto.
(* EF_builtin *)
- eapply do_external_function_complete; eauto.
+ eapply BF_EX; eauto.
(* EF_runtime *)
- eapply do_external_function_complete; eauto.
+ eapply BF_EX; eauto.
(* EF_vload *)
inv H; unfold do_ef_volatile_load.
exploit do_volatile_load_complete; eauto. intros EQ; rewrite EQ; auto.
@@ -1099,8 +1124,8 @@ Proof.
induction 1; intros; constructor; eauto.
Qed.
-Hint Constructors context contextlist.
-Hint Resolve context_compose contextlist_compose.
+Local Hint Constructors context contextlist : core.
+Local Hint Resolve context_compose contextlist_compose : core.
Definition reduction_ok (k: kind) (a: expr) (m: mem) (rd: reduction) : Prop :=
match k, rd with
@@ -1666,8 +1691,9 @@ Proof.
change (In (f (C0, rd)) (map f res2)). apply in_map; auto.
Qed.
-Hint Resolve reducts_incl_val reducts_incl_loc reducts_incl_listval
- reducts_incl_incontext reducts_incl_incontext2_left reducts_incl_incontext2_right.
+Local Hint Resolve reducts_incl_val reducts_incl_loc reducts_incl_listval
+ reducts_incl_incontext reducts_incl_incontext2_left
+ reducts_incl_incontext2_right : core.
Lemma step_expr_context:
forall from to C, context from to C ->
@@ -2052,7 +2078,7 @@ Ltac myinv :=
| _ => idtac
end.
-Hint Extern 3 => exact I.
+Local Hint Extern 3 => exact I : core.
Theorem do_step_sound:
forall w S rule t S',
diff --git a/cfrontend/Cop.v b/cfrontend/Cop.v
index c395a2cb..aa73abb0 100644
--- a/cfrontend/Cop.v
+++ b/cfrontend/Cop.v
@@ -1131,7 +1131,7 @@ Qed.
Remark val_inject_vptrofs: forall n, Val.inject f (Vptrofs n) (Vptrofs n).
Proof. intros. unfold Vptrofs. destruct Archi.ptr64; auto. Qed.
-Hint Resolve val_inject_vtrue val_inject_vfalse val_inject_of_bool val_inject_vptrofs.
+Local Hint Resolve val_inject_vtrue val_inject_vfalse val_inject_of_bool val_inject_vptrofs : core.
Ltac TrivialInject :=
match goal with
@@ -1517,7 +1517,7 @@ Inductive val_casted: val -> type -> Prop :=
| val_casted_void: forall v,
val_casted v Tvoid.
-Hint Constructors val_casted.
+Local Hint Constructors val_casted : core.
Remark cast_int_int_idem:
forall sz sg i, cast_int_int sz sg (cast_int_int sz sg i) = cast_int_int sz sg i.
@@ -1580,6 +1580,27 @@ Proof.
intros. apply cast_val_casted. eapply cast_val_is_casted; eauto.
Qed.
+(** Moreover, casted values belong to the machine type corresponding to the
+ C type. *)
+
+Lemma val_casted_has_type:
+ forall v ty, val_casted v ty -> ty <> Tvoid -> Val.has_type v (typ_of_type ty).
+Proof.
+ intros. inv H; simpl typ_of_type.
+- exact I.
+- exact I.
+- exact I.
+- exact I.
+- apply Val.Vptr_has_type.
+- red; unfold Tptr; rewrite H1; auto.
+- red; unfold Tptr; rewrite H1; auto.
+- red; unfold Tptr; rewrite H1; auto.
+- red; unfold Tptr; rewrite H1; auto.
+- apply Val.Vptr_has_type.
+- apply Val.Vptr_has_type.
+- congruence.
+Qed.
+
(** Relation with the arithmetic conversions of ISO C99, section 6.3.1 *)
Module ArithConv.
diff --git a/cfrontend/Csem.v b/cfrontend/Csem.v
index 0c3e00de..a76a14ba 100644
--- a/cfrontend/Csem.v
+++ b/cfrontend/Csem.v
@@ -15,19 +15,9 @@
(** Dynamic semantics for the Compcert C language *)
-Require Import Coqlib.
-Require Import Errors.
-Require Import Maps.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import AST.
-Require Import Memory.
-Require Import Events.
-Require Import Globalenvs.
-Require Import Ctypes.
-Require Import Cop.
-Require Import Csyntax.
+Require Import Coqlib Errors Maps.
+Require Import Integers Floats Values AST Memory Builtins Events Globalenvs.
+Require Import Ctypes Cop Csyntax.
Require Import Smallstep.
(** * Operational semantics *)
@@ -437,6 +427,59 @@ Definition not_stuck (e: expr) (m: mem) : Prop :=
forall k C e' ,
context k RV C -> e = C e' -> imm_safe k e' m.
+(** ** Derived forms. *)
+
+(** The following are admissible reduction rules for some derived forms
+ of the CompCert C language. They help showing that the derived forms
+ make sense. *)
+
+Lemma red_selection:
+ forall v1 ty1 v2 ty2 v3 ty3 ty m b v2' v3',
+ ty <> Tvoid ->
+ bool_val v1 ty1 m = Some b ->
+ sem_cast v2 ty2 ty m = Some v2' ->
+ sem_cast v3 ty3 ty m = Some v3' ->
+ rred (Eselection (Eval v1 ty1) (Eval v2 ty2) (Eval v3 ty3) ty) m
+ E0 (Eval (if b then v2' else v3') ty) m.
+Proof.
+ intros. unfold Eselection.
+ set (t := typ_of_type ty).
+ set (sg := mksignature (AST.Tint :: t :: t :: nil) (Some t) cc_default).
+ assert (LK: lookup_builtin_function "__builtin_sel"%string sg = Some (BI_standard (BI_select t))).
+ { unfold sg, t; destruct ty as [ | ? ? ? | ? | [] ? | ? ? | ? ? ? | ? ? ? | ? ? | ? ? ];
+ simpl; unfold Tptr; destruct Archi.ptr64; reflexivity. }
+ set (v' := if b then v2' else v3').
+ assert (C: val_casted v' ty).
+ { unfold v'; destruct b; eapply cast_val_is_casted; eauto. }
+ assert (EQ: Val.normalize v' t = v').
+ { apply Val.normalize_idem. apply val_casted_has_type; auto. }
+ econstructor.
+- constructor. rewrite cast_bool_bool_val, H0. eauto.
+ constructor. eauto.
+ constructor. eauto.
+ constructor.
+- red. red. rewrite LK. constructor. simpl. rewrite <- EQ.
+ destruct b; auto.
+Qed.
+
+Lemma ctx_selection_1:
+ forall k C r2 r3 ty, context k RV C -> context k RV (fun x => Eselection (C x) r2 r3 ty).
+Proof.
+ intros. apply ctx_builtin. constructor; auto.
+Qed.
+
+Lemma ctx_selection_2:
+ forall k r1 C r3 ty, context k RV C -> context k RV (fun x => Eselection r1 (C x) r3 ty).
+Proof.
+ intros. apply ctx_builtin. constructor; constructor; auto.
+Qed.
+
+Lemma ctx_selection_3:
+ forall k r1 r2 C ty, context k RV C -> context k RV (fun x => Eselection r1 r2 (C x) ty).
+Proof.
+ intros. apply ctx_builtin. constructor; constructor; constructor; auto.
+Qed.
+
End EXPR.
(** ** Transition semantics. *)
diff --git a/cfrontend/Cstrategy.v b/cfrontend/Cstrategy.v
index 28c8eeb8..c235031f 100644
--- a/cfrontend/Cstrategy.v
+++ b/cfrontend/Cstrategy.v
@@ -222,7 +222,7 @@ Proof.
induction 1; constructor; auto.
Qed.
-Hint Resolve leftcontext_context.
+Local Hint Resolve leftcontext_context : core.
(** Strategy for reducing expressions. We reduce the leftmost innermost
non-simple subexpression, evaluating its arguments (which are necessarily
@@ -398,8 +398,8 @@ Proof.
induction 1; intros; constructor; eauto.
Qed.
-Hint Constructors context contextlist.
-Hint Resolve context_compose contextlist_compose.
+Local Hint Constructors context contextlist : core.
+Local Hint Resolve context_compose contextlist_compose : core.
(** * Safe executions. *)
@@ -975,7 +975,7 @@ Proof.
apply extensionality; intros. f_equal. f_equal. apply exprlist_app_assoc.
Qed.
-Hint Resolve contextlist'_head contextlist'_tail.
+Local Hint Resolve contextlist'_head contextlist'_tail : core.
Lemma eval_simple_list_steps:
forall rl vl, eval_simple_list' rl vl ->
@@ -1049,7 +1049,7 @@ Scheme expr_ind2 := Induction for expr Sort Prop
with exprlist_ind2 := Induction for exprlist Sort Prop.
Combined Scheme expr_expr_list_ind from expr_ind2, exprlist_ind2.
-Hint Constructors leftcontext leftcontextlist.
+Local Hint Constructors leftcontext leftcontextlist : core.
Lemma decompose_expr:
(forall a from C,
diff --git a/cfrontend/Csyntax.v b/cfrontend/Csyntax.v
index 00565309..c34a5e13 100644
--- a/cfrontend/Csyntax.v
+++ b/cfrontend/Csyntax.v
@@ -100,6 +100,18 @@ Definition Epreincr (id: incr_or_decr) (l: expr) (ty: type) :=
Eassignop (match id with Incr => Oadd | Decr => Osub end)
l (Eval (Vint Int.one) type_int32s) (typeconv ty) ty.
+(** Selection is a conditional expression that always evaluates both arms
+ and can be implemented by "conditional move" instructions.
+ It is expressed as an invocation of a builtin function. *)
+
+Definition Eselection (r1 r2 r3: expr) (ty: type) :=
+ let t := typ_of_type ty in
+ let sg := mksignature (AST.Tint :: t :: t :: nil) (Some t) cc_default in
+ Ebuiltin (EF_builtin "__builtin_sel"%string sg)
+ (Tcons type_bool (Tcons ty (Tcons ty Tnil)))
+ (Econs r1 (Econs r2 (Econs r3 Enil)))
+ ty.
+
(** Extract the type part of a type-annotated expression. *)
Definition typeof (a: expr) : type :=
diff --git a/cfrontend/Ctyping.v b/cfrontend/Ctyping.v
index ba1d34fb..b92a9bac 100644
--- a/cfrontend/Ctyping.v
+++ b/cfrontend/Ctyping.v
@@ -18,7 +18,7 @@
Require Import String.
Require Import Coqlib Maps Integers Floats Errors.
Require Import AST Linking.
-Require Import Values Memory Globalenvs Events.
+Require Import Values Memory Globalenvs Builtins Events.
Require Import Ctypes Cop Csyntax Csem.
Local Open Scope error_monad_scope.
@@ -392,13 +392,17 @@ Inductive wt_rvalue : expr -> Prop :=
classify_fun (typeof r1) = fun_case_f tyargs tyres cconv ->
wt_arguments rargs tyargs ->
wt_rvalue (Ecall r1 rargs tyres)
- | wt_Ebuiltin: forall ef tyargs rargs,
+ | wt_Ebuiltin: forall ef tyargs rargs ty,
wt_exprlist rargs ->
wt_arguments rargs tyargs ->
- (* This is specialized to builtins returning void, the only
- case generated by C2C. *)
- sig_res (ef_sig ef) = None ->
- wt_rvalue (Ebuiltin ef tyargs rargs Tvoid)
+ (* This typing rule is specialized to the builtin invocations generated
+ by C2C, which are either __builtin_sel or builtins returning void. *)
+ (ty = Tvoid /\ sig_res (ef_sig ef) = None)
+ \/ (tyargs = Tcons type_bool (Tcons ty (Tcons ty Tnil))
+ /\ let t := typ_of_type ty in
+ let sg := mksignature (AST.Tint :: t :: t :: nil) (Some t) cc_default in
+ ef = EF_builtin "__builtin_sel"%string sg) ->
+ wt_rvalue (Ebuiltin ef tyargs rargs ty)
| wt_Eparen: forall r tycast ty,
wt_rvalue r ->
wt_cast (typeof r) tycast -> subtype tycast ty ->
@@ -745,6 +749,12 @@ Definition ebuiltin (ef: external_function) (tyargs: typelist) (args: exprlist)
then OK (Ebuiltin ef tyargs args tyres)
else Error (msg "builtin: wrong type decoration").
+Definition eselection (r1 r2 r3: expr) : res expr :=
+ do x1 <- check_rval r1; do x2 <- check_rval r2; do x3 <- check_rval r3;
+ do y1 <- check_bool (typeof r1);
+ do ty <- type_conditional (typeof r2) (typeof r3);
+ OK (Eselection r1 r2 r3 ty).
+
Definition sdo (a: expr) : res statement :=
do x <- check_rval a; OK (Sdo a).
@@ -981,6 +991,15 @@ Proof.
destruct i; auto.
Qed.
+Lemma wt_bool_cast:
+ forall ty, wt_bool ty -> wt_cast ty type_bool.
+Proof.
+ unfold wt_bool, wt_cast; unfold classify_bool; intros.
+ destruct ty; simpl in *; try congruence;
+ try (destruct Archi.ptr64; congruence).
+ destruct f; congruence.
+Qed.
+
Lemma wt_cast_int:
forall i1 s1 a1 i2 s2 a2, wt_cast (Tint i1 s1 a1) (Tint i2 s2 a2).
Proof.
@@ -1225,6 +1244,17 @@ Proof.
econstructor; eauto. eapply check_arguments_sound; eauto.
Qed.
+Lemma eselection_sound:
+ forall r1 r2 r3 a, eselection r1 r2 r3 = OK a ->
+ wt_expr ce e r1 -> wt_expr ce e r2 -> wt_expr ce e r3 -> wt_expr ce e a.
+Proof.
+ intros. monadInv H. apply type_conditional_cast in EQ3. destruct EQ3.
+ eapply wt_Ebuiltin.
+ repeat (constructor; eauto with ty).
+ repeat (constructor; eauto with ty). apply wt_bool_cast; eauto with ty.
+ right; auto.
+Qed.
+
Lemma sdo_sound:
forall a s, sdo a = OK s -> wt_expr ce e a -> wt_stmt ce e rt s.
Proof.
@@ -1632,15 +1662,6 @@ Proof.
destruct f; discriminate.
Qed.
-Lemma wt_bool_cast:
- forall ty, wt_bool ty -> wt_cast ty type_bool.
-Proof.
- unfold wt_bool, wt_cast; unfold classify_bool; intros.
- destruct ty; simpl in *; try congruence;
- try (destruct Archi.ptr64; congruence).
- destruct f; congruence.
-Qed.
-
Lemma wt_cast_self:
forall t1 t2, wt_cast t1 t2 -> wt_cast t2 t2.
Proof.
@@ -1725,7 +1746,27 @@ Proof.
constructor; auto.
- (* comma *) auto.
- (* paren *) inv H3. constructor. apply H5. eapply pres_sem_cast; eauto.
-- (* builtin *) subst. auto with ty.
+- (* builtin *) subst. destruct H7 as [(A & B) | (A & B)].
++ subst ty. auto with ty.
++ simpl in B. set (T := typ_of_type ty) in *.
+ set (sg := mksignature (AST.Tint :: T :: T :: nil) (Some T) cc_default) in *.
+ assert (LK: lookup_builtin_function "__builtin_sel"%string sg = Some (BI_standard (BI_select T))).
+ { unfold sg, T; destruct ty as [ | ? ? ? | ? | [] ? | ? ? | ? ? ? | ? ? ? | ? ? | ? ? ];
+ simpl; unfold Tptr; destruct Archi.ptr64; reflexivity. }
+ subst ef. red in H0. red in H0. rewrite LK in H0. inv H0.
+ inv H. inv H8. inv H9. inv H10. simpl in H1.
+ assert (A: val_casted v1 type_bool) by (eapply cast_val_is_casted; eauto).
+ inv A.
+ set (v' := if Int.eq n Int.zero then v4 else v2) in *.
+ constructor.
+ destruct (type_eq ty Tvoid).
+ subst. constructor.
+ inv H1.
+ assert (C: val_casted v' ty).
+ { unfold v'; destruct (Int.eq n Int.zero); eapply cast_val_is_casted; eauto. }
+ assert (EQ: Val.normalize v' T = v').
+ { apply Val.normalize_idem. apply val_casted_has_type; auto. }
+ rewrite EQ. apply wt_val_casted; auto.
Qed.
Lemma wt_lred:
@@ -1767,8 +1808,8 @@ with wt_subexprlist:
wt_exprlist cenv tenv (C a) ->
wt_expr_kind cenv tenv from a.
Proof.
- destruct 1; intros WT; auto; inv WT; eauto.
- destruct 1; intros WT; inv WT; eauto.
+- destruct 1; intros WT; auto; inv WT; eauto.
+- destruct 1; intros WT; inv WT; eauto.
Qed.
Lemma typeof_context:
diff --git a/cfrontend/SimplExprspec.v b/cfrontend/SimplExprspec.v
index 37e2cd96..e7d57a1c 100644
--- a/cfrontend/SimplExprspec.v
+++ b/cfrontend/SimplExprspec.v
@@ -687,7 +687,7 @@ Hint Resolve gensym_within within_widen contained_widen
in_eq in_cons
Ple_trans Ple_refl: gensym.
-Hint Resolve dest_for_val_below dest_for_effect_below.
+Local Hint Resolve dest_for_val_below dest_for_effect_below : core.
(** ** Correctness of the translation functions *)
diff --git a/common/Builtins.v b/common/Builtins.v
new file mode 100644
index 00000000..c9097e86
--- /dev/null
+++ b/common/Builtins.v
@@ -0,0 +1,58 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, Collège de France and Inria Paris *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the GNU General Public License as published by *)
+(* the Free Software Foundation, either version 2 of the License, or *)
+(* (at your option) any later version. This file is also distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Known built-in functions *)
+
+Require Import String Coqlib.
+Require Import AST Integers Floats Values.
+Require Export Builtins0 Builtins1.
+
+Inductive builtin_function : Type :=
+ | BI_standard (b: standard_builtin)
+ | BI_platform (b: platform_builtin).
+
+Definition builtin_function_sig (b: builtin_function) : signature :=
+ match b with
+ | BI_standard b => standard_builtin_sig b
+ | BI_platform b => platform_builtin_sig b
+ end.
+
+Definition builtin_function_sem (b: builtin_function) : builtin_sem (proj_sig_res (builtin_function_sig b)) :=
+ match b with
+ | BI_standard b => standard_builtin_sem b
+ | BI_platform b => platform_builtin_sem b
+ end.
+
+Definition lookup_builtin_function (name: string) (sg: signature) : option builtin_function :=
+ match lookup_builtin standard_builtin_sig name sg standard_builtin_table with
+ | Some b => Some (BI_standard b)
+ | None =>
+ match lookup_builtin platform_builtin_sig name sg platform_builtin_table with
+ | Some b => Some (BI_platform b)
+ | None => None
+ end end.
+
+Lemma lookup_builtin_function_sig:
+ forall name sg b, lookup_builtin_function name sg = Some b -> builtin_function_sig b = sg.
+Proof.
+ unfold lookup_builtin_function; intros.
+ destruct (lookup_builtin standard_builtin_sig name sg standard_builtin_table) as [bs|] eqn:E.
+ inv H. simpl. eapply lookup_builtin_sig; eauto.
+ destruct (lookup_builtin platform_builtin_sig name sg platform_builtin_table) as [bp|] eqn:E'.
+ inv H. simpl. eapply lookup_builtin_sig; eauto.
+ discriminate.
+Qed.
+
+
diff --git a/common/Builtins0.v b/common/Builtins0.v
new file mode 100644
index 00000000..b78006dd
--- /dev/null
+++ b/common/Builtins0.v
@@ -0,0 +1,531 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, Collège de France and Inria Paris *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the GNU General Public License as published by *)
+(* the Free Software Foundation, either version 2 of the License, or *)
+(* (at your option) any later version. This file is also distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Associating semantics to built-in functions *)
+
+Require Import String Coqlib.
+Require Import AST Integers Floats Values Memdata.
+
+(** This module provides definitions and mechanisms to associate semantics
+ with names of built-in functions.
+
+ This module is independent of the target architecture. Each target
+ provides a [Builtins1] module that lists the built-ins semantics
+ appropriate for the target.
+*)
+
+Definition val_opt_has_type (ov: option val) (t: typ) : Prop :=
+ match ov with Some v => Val.has_type v t | None => True end.
+
+Definition val_opt_inject (j: meminj) (ov ov': option val) : Prop :=
+ match ov, ov' with
+ | None, _ => True
+ | Some v, Some v' => Val.inject j v v'
+ | _, None => False
+ end.
+
+(** The semantics of a built-in function is a partial function
+ from list of values to values.
+ It must agree with the return type stated in the signature,
+ and be compatible with value injections.
+*)
+
+Record builtin_sem (tret: typ) : Type := mkbuiltin {
+ bs_sem :> list val -> option val;
+ bs_well_typed: forall vl,
+ val_opt_has_type (bs_sem vl) tret;
+ bs_inject: forall j vl vl',
+ Val.inject_list j vl vl' -> val_opt_inject j (bs_sem vl) (bs_sem vl')
+}.
+
+(** Builtin functions can be created from functions over values, such as those
+ from the [Values.Val] module. Proofs of well-typedness and compatibility with
+ injections must be provided. The constructor functions have names
+- [mkbuiltin_vNt] for a [t]otal function of [N] arguments that are [v]alues, or
+- [mkbuiltin_vNp] for a [p]artial function of [N] arguments that are [v]alues.
+*)
+
+Local Unset Program Cases.
+
+Program Definition mkbuiltin_v1t
+ (tret: typ) (f: val -> val)
+ (WT: forall v1, Val.has_type (f v1) tret)
+ (INJ: forall j v1 v1', Val.inject j v1 v1' -> Val.inject j (f v1) (f v1')) :=
+ mkbuiltin tret (fun vl => match vl with v1 :: nil => Some (f v1) | _ => None end) _ _.
+Next Obligation.
+ red; destruct vl; auto. destruct vl; auto.
+Qed.
+Next Obligation.
+ red; inv H; auto. inv H1; auto.
+Qed.
+
+Program Definition mkbuiltin_v2t
+ (tret: typ) (f: val -> val -> val)
+ (WT: forall v1 v2, Val.has_type (f v1 v2) tret)
+ (INJ: forall j v1 v1' v2 v2',
+ Val.inject j v1 v1' -> Val.inject j v2 v2' ->
+ Val.inject j (f v1 v2) (f v1' v2')) :=
+ mkbuiltin tret (fun vl => match vl with v1 :: v2 :: nil => Some (f v1 v2) | _ => None end) _ _.
+Next Obligation.
+ red; destruct vl; auto. destruct vl; auto. destruct vl; auto.
+Qed.
+Next Obligation.
+ red; inv H; auto. inv H1; auto. inv H2; auto.
+Qed.
+
+Program Definition mkbuiltin_v3t
+ (tret: typ) (f: val -> val -> val -> val)
+ (WT: forall v1 v2 v3, Val.has_type (f v1 v2 v3) tret)
+ (INJ: forall j v1 v1' v2 v2' v3 v3',
+ Val.inject j v1 v1' -> Val.inject j v2 v2' -> Val.inject j v3 v3' ->
+ Val.inject j (f v1 v2 v3) (f v1' v2' v3')) :=
+ mkbuiltin tret (fun vl => match vl with v1 :: v2 :: v3 :: nil => Some (f v1 v2 v3) | _ => None end) _ _.
+Next Obligation.
+ red; destruct vl; auto. destruct vl; auto. destruct vl; auto. destruct vl; auto.
+Qed.
+Next Obligation.
+ red; inv H; auto. inv H1; auto. inv H2; auto. inv H3; auto.
+Qed.
+
+Program Definition mkbuiltin_v1p
+ (tret: typ) (f: val -> option val)
+ (WT: forall v1, val_opt_has_type (f v1) tret)
+ (INJ: forall j v1 v1',
+ Val.inject j v1 v1' -> val_opt_inject j (f v1) (f v1')) :=
+ mkbuiltin tret (fun vl => match vl with v1 :: nil => f v1 | _ => None end) _ _.
+Next Obligation.
+ red; destruct vl; auto. destruct vl; auto. apply WT.
+Qed.
+Next Obligation.
+ red; inv H; auto. inv H1; auto. apply INJ; auto.
+Qed.
+
+Program Definition mkbuiltin_v2p
+ (tret: typ) (f: val -> val -> option val)
+ (WT: forall v1 v2, val_opt_has_type (f v1 v2) tret)
+ (INJ: forall j v1 v1' v2 v2',
+ Val.inject j v1 v1' -> Val.inject j v2 v2' ->
+ val_opt_inject j (f v1 v2) (f v1' v2')) :=
+ mkbuiltin tret (fun vl => match vl with v1 :: v2 :: nil => f v1 v2 | _ => None end) _ _.
+Next Obligation.
+ red; destruct vl; auto. destruct vl; auto. destruct vl; auto. apply WT.
+Qed.
+Next Obligation.
+ red; inv H; auto. inv H1; auto. inv H2; auto. apply INJ; auto.
+Qed.
+
+(** For numerical functions, involving only integers and floating-point numbers
+ but no pointer values, we can automate the proofs of well-typedness and
+ of compatibility with injections. *)
+
+(** First we define a mapping from syntactic Cminor types ([Tint], [Tfloat], etc) to semantic Coq types. *)
+
+Definition valty (t: typ) : Type :=
+ match t with
+ | Tint => int
+ | Tlong => int64
+ | Tfloat => float
+ | Tsingle => float32
+ | Tany32 | Tany64 => Empty_set (**r no clear semantic meaning *)
+ end.
+
+(** We can inject and project between the numerical type [valty t] and the type [val]. *)
+
+Definition inj_num (t: typ) : valty t -> val :=
+ match t with
+ | Tint => Vint
+ | Tlong => Vlong
+ | Tfloat => Vfloat
+ | Tsingle => Vsingle
+ | Tany32 | Tany64 => fun _ => Vundef
+ end.
+
+Definition proj_num {A: Type} (t: typ) (k0: A) (v: val): (valty t -> A) -> A :=
+ match t with
+ | Tint => fun k1 => match v with Vint n => k1 n | _ => k0 end
+ | Tlong => fun k1 => match v with Vlong n => k1 n | _ => k0 end
+ | Tfloat => fun k1 => match v with Vfloat n => k1 n | _ => k0 end
+ | Tsingle => fun k1 => match v with Vsingle n => k1 n | _ => k0 end
+ | Tany32 | Tany64 => fun k1 => k0
+ end.
+
+Lemma inj_num_wt: forall t x, Val.has_type (inj_num t x) t.
+Proof.
+ destruct t; intros; exact I.
+Qed.
+
+Lemma inj_num_inject: forall j t x, Val.inject j (inj_num t x) (inj_num t x).
+Proof.
+ destruct t; intros; constructor.
+Qed.
+
+Lemma inj_num_opt_wt: forall t x, val_opt_has_type (option_map (inj_num t) x) t.
+Proof.
+ intros. destruct x; simpl. apply inj_num_wt. auto.
+Qed.
+
+Lemma inj_num_opt_inject: forall j t x,
+ val_opt_inject j (option_map (inj_num t) x) (option_map (inj_num t) x).
+Proof.
+ destruct x; simpl. apply inj_num_inject. auto.
+Qed.
+
+Lemma proj_num_wt:
+ forall tres t k1 v,
+ (forall x, Val.has_type (k1 x) tres) ->
+ Val.has_type (proj_num t Vundef v k1) tres.
+Proof.
+ intros. destruct t; simpl; destruct v; auto; exact I.
+Qed.
+
+Lemma proj_num_inject:
+ forall j t k1 v k1' v',
+ (forall x, Val.inject j (k1 x) (k1' x)) ->
+ Val.inject j v v' ->
+ Val.inject j (proj_num t Vundef v k1) (proj_num t Vundef v' k1').
+Proof.
+ intros. destruct t; simpl; inv H0; auto.
+Qed.
+
+Lemma proj_num_opt_wt:
+ forall tres t k0 k1 v,
+ k0 = None \/ k0 = Some Vundef ->
+ (forall x, val_opt_has_type (k1 x) tres) ->
+ val_opt_has_type (proj_num t k0 v k1) tres.
+Proof.
+ intros.
+ assert (val_opt_has_type k0 tres). { destruct H; subst k0; exact I. }
+ destruct t; simpl; destruct v; auto.
+Qed.
+
+Lemma proj_num_opt_inject:
+ forall j k0 t k1 v k1' v',
+ (forall ov, val_opt_inject j k0 ov) ->
+ (forall x, val_opt_inject j (k1 x) (k1' x)) ->
+ Val.inject j v v' ->
+ val_opt_inject j (proj_num t k0 v k1) (proj_num t k0 v' k1').
+Proof.
+ intros. destruct t; simpl; inv H1; auto.
+Qed.
+
+(** Wrapping numerical functions as built-ins. The constructor functions
+ have names
+- [mkbuiltin_nNt] for a [t]otal function of [N] numbers, or
+- [mkbuiltin_vNp] for a [p]artial function of [N] numbers.
+ *)
+
+Program Definition mkbuiltin_n1t
+ (targ1: typ) (tres: typ)
+ (f: valty targ1 -> valty tres) :=
+ mkbuiltin_v1t tres
+ (fun v1 => proj_num targ1 Vundef v1 (fun x => inj_num tres (f x)))
+ _ _.
+Next Obligation.
+ auto using proj_num_wt, inj_num_wt.
+Qed.
+Next Obligation.
+ auto using proj_num_inject, inj_num_inject.
+Qed.
+
+Program Definition mkbuiltin_n2t
+ (targ1 targ2: typ) (tres: typ)
+ (f: valty targ1 -> valty targ2 -> valty tres) :=
+ mkbuiltin_v2t tres
+ (fun v1 v2 =>
+ proj_num targ1 Vundef v1 (fun x1 =>
+ proj_num targ2 Vundef v2 (fun x2 => inj_num tres (f x1 x2))))
+ _ _.
+Next Obligation.
+ auto using proj_num_wt, inj_num_wt.
+Qed.
+Next Obligation.
+ auto using proj_num_inject, inj_num_inject.
+Qed.
+
+Program Definition mkbuiltin_n3t
+ (targ1 targ2 targ3: typ) (tres: typ)
+ (f: valty targ1 -> valty targ2 -> valty targ3 -> valty tres) :=
+ mkbuiltin_v3t tres
+ (fun v1 v2 v3 =>
+ proj_num targ1 Vundef v1 (fun x1 =>
+ proj_num targ2 Vundef v2 (fun x2 =>
+ proj_num targ3 Vundef v3 (fun x3 => inj_num tres (f x1 x2 x3)))))
+ _ _.
+Next Obligation.
+ auto using proj_num_wt, inj_num_wt.
+Qed.
+Next Obligation.
+ auto using proj_num_inject, inj_num_inject.
+Qed.
+
+Program Definition mkbuiltin_n1p
+ (targ1: typ) (tres: typ)
+ (f: valty targ1 -> option (valty tres)) :=
+ mkbuiltin_v1p tres
+ (fun v1 => proj_num targ1 None v1 (fun x => option_map (inj_num tres) (f x)))
+ _ _.
+Next Obligation.
+ auto using proj_num_opt_wt, inj_num_opt_wt.
+Qed.
+Next Obligation.
+ apply proj_num_opt_inject; auto. intros; red; auto. intros; apply inj_num_opt_inject.
+Qed.
+
+Program Definition mkbuiltin_n2p
+ (targ1 targ2: typ) (tres: typ)
+ (f: valty targ1 -> valty targ2 -> option (valty tres)) :=
+ mkbuiltin_v2p tres
+ (fun v1 v2 =>
+ proj_num targ1 None v1 (fun x1 =>
+ proj_num targ2 None v2 (fun x2 => option_map (inj_num tres) (f x1 x2))))
+ _ _.
+Next Obligation.
+ auto using proj_num_opt_wt, inj_num_opt_wt.
+Qed.
+Next Obligation.
+ apply proj_num_opt_inject; auto. intros; red; auto. intros.
+ apply proj_num_opt_inject; auto. intros; red; auto. intros.
+ apply inj_num_opt_inject.
+Qed.
+
+(** Looking up builtins by name and signature *)
+
+Section LOOKUP.
+
+Context {A: Type} (sig_of: A -> signature).
+
+Fixpoint lookup_builtin (name: string) (sg: signature) (l: list (string * A)) : option A :=
+ match l with
+ | nil => None
+ | (n, b) :: l =>
+ if string_dec name n && signature_eq sg (sig_of b)
+ then Some b
+ else lookup_builtin name sg l
+ end.
+
+Lemma lookup_builtin_sig: forall name sg b l,
+ lookup_builtin name sg l = Some b -> sig_of b = sg.
+Proof.
+ induction l as [ | [n b'] l ]; simpl; intros.
+- discriminate.
+- destruct (string_dec name n && signature_eq sg (sig_of b')) eqn:E.
++ InvBooleans. congruence.
++ auto.
+Qed.
+
+End LOOKUP.
+
+(** The standard, platform-independent built-ins *)
+
+Inductive standard_builtin : Type :=
+ | BI_select (t: typ)
+ | BI_fabs
+ | BI_fsqrt
+ | BI_negl
+ | BI_addl
+ | BI_subl
+ | BI_mull
+ | BI_i16_bswap
+ | BI_i32_bswap
+ | BI_i64_bswap
+ | BI_i64_umulh
+ | BI_i64_smulh
+ | BI_i64_sdiv
+ | BI_i64_udiv
+ | BI_i64_smod
+ | BI_i64_umod
+ | BI_i64_shl
+ | BI_i64_shr
+ | BI_i64_sar
+ | BI_i64_dtos
+ | BI_i64_dtou
+ | BI_i64_stod
+ | BI_i64_utod
+ | BI_i64_stof
+ | BI_i64_utof.
+
+Local Open Scope string_scope.
+
+Definition standard_builtin_table : list (string * standard_builtin) :=
+ ("__builtin_sel", BI_select Tint)
+ :: ("__builtin_sel", BI_select Tlong)
+ :: ("__builtin_sel", BI_select Tfloat)
+ :: ("__builtin_sel", BI_select Tsingle)
+ :: ("__builtin_fabs", BI_fabs)
+ :: ("__builtin_fsqrt", BI_fsqrt)
+ :: ("__builtin_negl", BI_negl)
+ :: ("__builtin_addl", BI_addl)
+ :: ("__builtin_subl", BI_subl)
+ :: ("__builtin_mull", BI_mull)
+ :: ("__builtin_bswap16", BI_i16_bswap)
+ :: ("__builtin_bswap", BI_i32_bswap)
+ :: ("__builtin_bswap32", BI_i32_bswap)
+ :: ("__builtin_bswap64", BI_i64_bswap)
+ :: ("__compcert_i64_umulh", BI_i64_umulh)
+ :: ("__compcert_i64_smulh", BI_i64_smulh)
+ :: ("__compcert_i64_sdiv", BI_i64_sdiv)
+ :: ("__compcert_i64_udiv", BI_i64_udiv)
+ :: ("__compcert_i64_smod", BI_i64_smod)
+ :: ("__compcert_i64_umod", BI_i64_umod)
+ :: ("__compcert_i64_shl", BI_i64_shl)
+ :: ("__compcert_i64_shr", BI_i64_shr)
+ :: ("__compcert_i64_sar", BI_i64_sar)
+ :: ("__compcert_i64_dtos", BI_i64_dtos)
+ :: ("__compcert_i64_dtou", BI_i64_dtou)
+ :: ("__compcert_i64_stod", BI_i64_stod)
+ :: ("__compcert_i64_utod", BI_i64_utod)
+ :: ("__compcert_i64_stof", BI_i64_stof)
+ :: ("__compcert_i64_utof", BI_i64_utof)
+ :: nil.
+
+Definition standard_builtin_sig (b: standard_builtin) : signature :=
+ match b with
+ | BI_select t =>
+ mksignature (Tint :: t :: t :: nil) (Some t) cc_default
+ | BI_fabs | BI_fsqrt =>
+ mksignature (Tfloat :: nil) (Some Tfloat) cc_default
+ | BI_negl =>
+ mksignature (Tlong :: nil) (Some Tlong) cc_default
+ | BI_addl | BI_subl | BI_i64_umulh| BI_i64_smulh
+ | BI_i64_sdiv | BI_i64_udiv | BI_i64_smod | BI_i64_umod =>
+ mksignature (Tlong :: Tlong :: nil) (Some Tlong) cc_default
+ | BI_mull =>
+ mksignature (Tint :: Tint :: nil) (Some Tlong) cc_default
+ | BI_i32_bswap =>
+ mksignature (Tint :: nil) (Some Tint) cc_default
+ | BI_i64_bswap =>
+ mksignature (Tlong :: nil) (Some Tlong) cc_default
+ | BI_i16_bswap =>
+ mksignature (Tint :: nil) (Some Tint) cc_default
+ | BI_i64_shl | BI_i64_shr | BI_i64_sar =>
+ mksignature (Tlong :: Tint :: nil) (Some Tlong) cc_default
+ | BI_i64_dtos | BI_i64_dtou =>
+ mksignature (Tfloat :: nil) (Some Tlong) cc_default
+ | BI_i64_stod | BI_i64_utod =>
+ mksignature (Tlong :: nil) (Some Tfloat) cc_default
+ | BI_i64_stof | BI_i64_utof =>
+ mksignature (Tlong :: nil) (Some Tsingle) cc_default
+ end.
+
+Program Definition standard_builtin_sem (b: standard_builtin) : builtin_sem (proj_sig_res (standard_builtin_sig b)) :=
+ match b with
+ | BI_select t =>
+ mkbuiltin t
+ (fun vargs => match vargs with
+ | Vint n :: v1 :: v2 :: nil => Some (Val.normalize (if Int.eq n Int.zero then v2 else v1) t)
+ | _ => None
+ end) _ _
+ | BI_fabs => mkbuiltin_n1t Tfloat Tfloat Float.abs
+ | BI_fsqrt => mkbuiltin_n1t Tfloat Tfloat Float.sqrt
+ | BI_negl => mkbuiltin_n1t Tlong Tlong Int64.neg
+ | BI_addl => mkbuiltin_v2t Tlong Val.addl _ _
+ | BI_subl => mkbuiltin_v2t Tlong Val.subl _ _
+ | BI_mull => mkbuiltin_v2t Tlong Val.mull' _ _
+ | BI_i16_bswap =>
+ mkbuiltin_n1t Tint Tint
+ (fun n => Int.repr (decode_int (List.rev (encode_int 2%nat (Int.unsigned n)))))
+ | BI_i32_bswap =>
+ mkbuiltin_n1t Tint Tint
+ (fun n => Int.repr (decode_int (List.rev (encode_int 4%nat (Int.unsigned n)))))
+ | BI_i64_bswap =>
+ mkbuiltin_n1t Tlong Tlong
+ (fun n => Int64.repr (decode_int (List.rev (encode_int 8%nat (Int64.unsigned n)))))
+ | BI_i64_umulh => mkbuiltin_n2t Tlong Tlong Tlong Int64.mulhu
+ | BI_i64_smulh => mkbuiltin_n2t Tlong Tlong Tlong Int64.mulhs
+ | BI_i64_sdiv => mkbuiltin_v2p Tlong Val.divls _ _
+ | BI_i64_udiv => mkbuiltin_v2p Tlong Val.divlu _ _
+ | BI_i64_smod => mkbuiltin_v2p Tlong Val.modls _ _
+ | BI_i64_umod => mkbuiltin_v2p Tlong Val.modlu _ _
+ | BI_i64_shl => mkbuiltin_v2t Tlong Val.shll _ _
+ | BI_i64_shr => mkbuiltin_v2t Tlong Val.shrlu _ _
+ | BI_i64_sar => mkbuiltin_v2t Tlong Val.shrl _ _
+ | BI_i64_dtos => mkbuiltin_n1p Tfloat Tlong Float.to_long
+ | BI_i64_dtou => mkbuiltin_n1p Tfloat Tlong Float.to_longu
+ | BI_i64_stod => mkbuiltin_n1t Tlong Tfloat Float.of_long
+ | BI_i64_utod => mkbuiltin_n1t Tlong Tfloat Float.of_longu
+ | BI_i64_stof => mkbuiltin_n1t Tlong Tsingle Float32.of_long
+ | BI_i64_utof => mkbuiltin_n1t Tlong Tsingle Float32.of_longu
+ end.
+Next Obligation.
+ red. destruct vl; auto. destruct v; auto.
+ destruct vl; auto. destruct vl; auto. destruct vl; auto.
+ apply Val.normalize_type.
+Qed.
+Next Obligation.
+ red. inv H; auto. inv H0; auto. inv H1; auto. inv H0; auto. inv H2; auto.
+ apply Val.normalize_inject. destruct (Int.eq i Int.zero); auto.
+Qed.
+Next Obligation.
+ unfold Val.addl, Val.has_type; destruct v1; auto; destruct v2; auto; destruct Archi.ptr64; auto.
+Qed.
+Next Obligation.
+ apply Val.addl_inject; auto.
+Qed.
+Next Obligation.
+ unfold Val.subl, Val.has_type, negb; destruct v1; auto; destruct v2; auto;
+ destruct Archi.ptr64; auto; destruct (eq_block b0 b1); auto.
+Qed.
+Next Obligation.
+ apply Val.subl_inject; auto.
+Qed.
+Next Obligation.
+ unfold Val.mull', Val.has_type; destruct v1; simpl; auto; destruct v2; auto.
+Qed.
+Next Obligation.
+ inv H; simpl; auto. inv H0; auto.
+Qed.
+Next Obligation.
+ red. destruct v1; simpl; auto. destruct v2; auto. destruct orb; exact I.
+Qed.
+Next Obligation.
+ red. inv H; simpl; auto. inv H0; auto. destruct orb; auto.
+Qed.
+Next Obligation.
+ red. destruct v1; simpl; auto. destruct v2; auto. destruct Int64.eq; exact I.
+Qed.
+Next Obligation.
+ red. inv H; simpl; auto. inv H0; auto. destruct Int64.eq; auto.
+Qed.
+Next Obligation.
+ red. destruct v1; simpl; auto. destruct v2; auto. destruct orb; exact I.
+Qed.
+Next Obligation.
+ red. inv H; simpl; auto. inv H0; auto. destruct orb; auto.
+Qed.
+Next Obligation.
+ red. destruct v1; simpl; auto. destruct v2; auto. destruct Int64.eq; exact I.
+Qed.
+Next Obligation.
+ red. inv H; simpl; auto. inv H0; auto. destruct Int64.eq; auto.
+Qed.
+Next Obligation.
+ red. destruct v1; simpl; auto. destruct v2; auto. destruct Int.ltu; auto.
+Qed.
+Next Obligation.
+ inv H; simpl; auto. inv H0; auto. destruct Int.ltu; auto.
+Qed.
+Next Obligation.
+ red. destruct v1; simpl; auto. destruct v2; auto. destruct Int.ltu; auto.
+Qed.
+Next Obligation.
+ inv H; simpl; auto. inv H0; auto. destruct Int.ltu; auto.
+Qed.
+Next Obligation.
+ red. destruct v1; simpl; auto. destruct v2; auto. destruct Int.ltu; auto.
+Qed.
+Next Obligation.
+ inv H; simpl; auto. inv H0; auto. destruct Int.ltu; auto.
+Qed.
+
diff --git a/common/Events.v b/common/Events.v
index 26dd505f..3fb84f49 100644
--- a/common/Events.v
+++ b/common/Events.v
@@ -24,6 +24,7 @@ Require Import Floats.
Require Import Values.
Require Import Memory.
Require Import Globalenvs.
+Require Import Builtins.
(** * Events and traces *)
@@ -1377,12 +1378,67 @@ Proof.
split. constructor. auto.
Qed.
+(** ** Semantics of known built-in functions. *)
+
+(** Some built-in functions and runtime support functions have known semantics
+ as defined in the [Builtin] modules.
+ These built-in functions have no observable effects and do not access memory. *)
+
+Inductive known_builtin_sem (bf: builtin_function) (ge: Senv.t):
+ list val -> mem -> trace -> val -> mem -> Prop :=
+ | known_builtin_sem_intro: forall vargs vres m,
+ builtin_function_sem bf vargs = Some vres ->
+ known_builtin_sem bf ge vargs m E0 vres m.
+
+Lemma known_builtin_ok: forall bf,
+ extcall_properties (known_builtin_sem bf) (builtin_function_sig bf).
+Proof.
+ intros. set (bsem := builtin_function_sem bf). constructor; intros.
+(* well typed *)
+- inv H.
+ specialize (bs_well_typed _ bsem vargs). unfold val_opt_has_type, bsem; rewrite H0.
+ auto.
+(* symbols *)
+- inv H0. econstructor; eauto.
+(* valid blocks *)
+- inv H; auto.
+(* perms *)
+- inv H; auto.
+(* readonly *)
+- inv H. apply Mem.unchanged_on_refl.
+(* mem extends *)
+- inv H. fold bsem in H2. apply val_inject_list_lessdef in H1.
+ specialize (bs_inject _ bsem _ _ _ H1).
+ unfold val_opt_inject; rewrite H2; intros.
+ destruct (bsem vargs') as [vres'|] eqn:?; try contradiction.
+ exists vres', m1'; intuition auto using Mem.extends_refl, Mem.unchanged_on_refl.
+ constructor; auto.
+ apply val_inject_lessdef; auto.
+(* mem injects *)
+- inv H0. fold bsem in H3.
+ specialize (bs_inject _ bsem _ _ _ H2).
+ unfold val_opt_inject; rewrite H3; intros.
+ destruct (bsem vargs') as [vres'|] eqn:?; try contradiction.
+ exists f, vres', m1'; intuition auto using Mem.extends_refl, Mem.unchanged_on_refl.
+ constructor; auto.
+ red; intros; congruence.
+(* trace length *)
+- inv H; simpl; omega.
+(* receptive *)
+- inv H; inv H0. exists vres1, m1; constructor; auto.
+(* determ *)
+- inv H; inv H0.
+ split. constructor. intuition congruence.
+Qed.
+
(** ** Semantics of external functions. *)
-(** For functions defined outside the program ([EF_external],
- [EF_builtin] and [EF_runtime]), we do not define their
- semantics, but only assume that it satisfies
- [extcall_properties]. *)
+(** For functions defined outside the program ([EF_external]),
+ we do not define their semantics, but only assume that it satisfies
+ [extcall_properties].
+ We do the same for built-in functions and runtime support functions that
+ are not described in [Builtins].
+*)
Parameter external_functions_sem: String.string -> signature -> extcall_sem.
@@ -1398,6 +1454,22 @@ Axiom inline_assembly_properties:
(** ** Combined semantics of external calls *)
+Definition builtin_or_external_sem name sg :=
+ match lookup_builtin_function name sg with
+ | Some bf => known_builtin_sem bf
+ | None => external_functions_sem name sg
+ end.
+
+Lemma builtin_or_external_sem_ok: forall name sg,
+ extcall_properties (builtin_or_external_sem name sg) sg.
+Proof.
+ unfold builtin_or_external_sem; intros.
+ destruct (lookup_builtin_function name sg) as [bf|] eqn:L.
+- exploit lookup_builtin_function_sig; eauto. intros EQ; subst sg.
+ apply known_builtin_ok.
+- apply external_functions_properties.
+Qed.
+
(** Combining the semantics given above for the various kinds of external calls,
we define the predicate [external_call] that relates:
- the external function being invoked
@@ -1412,8 +1484,8 @@ This predicate is used in the semantics of all CompCert languages. *)
Definition external_call (ef: external_function): extcall_sem :=
match ef with
| EF_external name sg => external_functions_sem name sg
- | EF_builtin name sg => external_functions_sem name sg
- | EF_runtime name sg => external_functions_sem name sg
+ | EF_builtin name sg => builtin_or_external_sem name sg
+ | EF_runtime name sg => builtin_or_external_sem name sg
| EF_vload chunk => volatile_load_sem chunk
| EF_vstore chunk => volatile_store_sem chunk
| EF_malloc => extcall_malloc_sem
@@ -1431,8 +1503,8 @@ Theorem external_call_spec:
Proof.
intros. unfold external_call, ef_sig; destruct ef.
apply external_functions_properties.
- apply external_functions_properties.
- apply external_functions_properties.
+ apply builtin_or_external_sem_ok.
+ apply builtin_or_external_sem_ok.
apply volatile_load_ok.
apply volatile_store_ok.
apply extcall_malloc_ok.
diff --git a/common/Separation.v b/common/Separation.v
index 1493b535..27065d1f 100644
--- a/common/Separation.v
+++ b/common/Separation.v
@@ -113,7 +113,7 @@ Proof.
intros P Q [[A B] [C D]]. split; auto.
Qed.
-Hint Resolve massert_imp_refl massert_eqv_refl.
+Hint Resolve massert_imp_refl massert_eqv_refl : core.
(** * Separating conjunction *)
diff --git a/common/Values.v b/common/Values.v
index e4297183..59e6388f 100644
--- a/common/Values.v
+++ b/common/Values.v
@@ -1950,7 +1950,7 @@ Inductive lessdef_list: list val -> list val -> Prop :=
lessdef v1 v2 -> lessdef_list vl1 vl2 ->
lessdef_list (v1 :: vl1) (v2 :: vl2).
-Hint Resolve lessdef_refl lessdef_undef lessdef_list_nil lessdef_list_cons.
+Hint Resolve lessdef_refl lessdef_undef lessdef_list_nil lessdef_list_cons : core.
Lemma lessdef_list_inv:
forall vl1 vl2, lessdef_list vl1 vl2 -> vl1 = vl2 \/ In Vundef vl1.
@@ -2175,7 +2175,7 @@ Inductive inject (mi: meminj): val -> val -> Prop :=
| val_inject_undef: forall v,
inject mi Vundef v.
-Hint Constructors inject.
+Hint Constructors inject : core.
Inductive inject_list (mi: meminj): list val -> list val-> Prop:=
| inject_list_nil :
@@ -2184,7 +2184,7 @@ Inductive inject_list (mi: meminj): list val -> list val-> Prop:=
inject mi v v' -> inject_list mi vl vl'->
inject_list mi (v :: vl) (v' :: vl').
-Hint Resolve inject_list_nil inject_list_cons.
+Hint Resolve inject_list_nil inject_list_cons : core.
Lemma inject_ptrofs:
forall mi i, inject mi (Vptrofs i) (Vptrofs i).
@@ -2192,7 +2192,7 @@ Proof.
unfold Vptrofs; intros. destruct Archi.ptr64; auto.
Qed.
-Hint Resolve inject_ptrofs.
+Hint Resolve inject_ptrofs : core.
Section VAL_INJ_OPS.
@@ -2495,7 +2495,7 @@ Proof.
constructor. eapply val_inject_incr; eauto. auto.
Qed.
-Hint Resolve inject_incr_refl val_inject_incr val_inject_list_incr.
+Hint Resolve inject_incr_refl val_inject_incr val_inject_list_incr : core.
Lemma val_inject_lessdef:
forall v1 v2, Val.lessdef v1 v2 <-> Val.inject (fun b => Some(b, 0)) v1 v2.
diff --git a/configure b/configure
index 00557673..3bdb5d2b 100755
--- a/configure
+++ b/configure
@@ -541,14 +541,14 @@ missingtools=false
echo "Testing Coq... " | tr -d '\n'
coq_ver=$(${COQBIN}coqc -v 2>/dev/null | sed -n -e 's/The Coq Proof Assistant, version \([^ ]*\).*$/\1/p')
case "$coq_ver" in
- 8.7.0|8.7.1|8.7.2|8.8.0|8.8.1|8.8.2|8.9.0|8.9.1|8.10+beta2)
+ 8.7.0|8.7.1|8.7.2|8.8.0|8.8.1|8.8.2|8.9.0|8.9.1|8.10)
echo "version $coq_ver -- good!";;
?*)
echo "version $coq_ver -- UNSUPPORTED"
if $ignore_coq_version; then
echo "Warning: this version of Coq is unsupported, proceed at your own risks."
else
- echo "Error: CompCert requires one of the following Coq versions: 8.9.1, 8.9.0, 8.8.2, 8.8.1, 8.8.0, 8.7.2, 8.7.1, 8.7.0"
+ echo "Error: CompCert requires one of the following Coq versions: 8.10, 8.9.1, 8.9.0, 8.8.2, 8.8.1, 8.8.0, 8.7.2, 8.7.1, 8.7.0"
missingtools=true
fi;;
"")
@@ -591,29 +591,23 @@ else
ocaml_opt_comp=false
fi
-MENHIR_REQUIRED=20161201
-MENHIR_NEW_API=20180530
-MENHIR_MAX=20181113
-menhir_flags=''
+MENHIR_REQUIRED=20190626
echo "Testing Menhir... " | tr -d '\n'
menhir_ver=`menhir --version 2>/dev/null | sed -n -e 's/^.*version \([0-9]*\).*$/\1/p'`
case "$menhir_ver" in
20[0-9][0-9][0-9][0-9][0-9][0-9])
- if test "$menhir_ver" -ge $MENHIR_REQUIRED -a "$menhir_ver" -le $MENHIR_MAX; then
+ if test "$menhir_ver" -ge $MENHIR_REQUIRED; then
echo "version $menhir_ver -- good!"
- menhir_include_dir=`menhir --suggest-menhirLib`
+ menhir_include_dir=$(menhir --suggest-menhirLib | tr -d '\r' | tr '\\' '/')
if test -z "$menhir_include_dir"; then
echo "Error: cannot determine the location of the Menhir API library."
echo "This can be due to an incorrect Menhir package."
echo "Consider using the OPAM package for Menhir."
missingtools=true
fi
- if test "$menhir_ver" -ge $MENHIR_NEW_API; then
- menhir_flags="--coq-lib-path compcert.cparser.MenhirLib"
- fi
else
echo "version $menhir_ver -- UNSUPPORTED"
- echo "Error: CompCert requires a version of Menhir between $MENHIR_REQUIRED and $MENHIR_MAX, included."
+ echo "Error: CompCert requires a version greater or equal to $MENHIR_REQUIRED."
missingtools=true
fi;;
*)
@@ -677,7 +671,8 @@ echo "-R lib compcert.lib \
-R driver compcert.driver \
-R flocq compcert.flocq \
-R exportclight compcert.exportclight \
--R cparser compcert.cparser" > _CoqProject
+-R cparser compcert.cparser \
+-R MenhirLib compcert.MenhirLib" > _CoqProject
case $arch in
x86)
echo "-R x86_${bitsize} compcert.x86_${bitsize}" >> _CoqProject
@@ -699,7 +694,6 @@ SHAREDIR=$sharedir
COQDEVDIR=$coqdevdir
OCAML_OPT_COMP=$ocaml_opt_comp
MENHIR_INCLUDES=-I "$menhir_include_dir"
-MENHIR_FLAGS=$menhir_flags
COMPFLAGS=-bin-annot
EOF
@@ -827,7 +821,7 @@ CFLAGS= -D __K1C_COS__
BACKENDLIB=Machblock.v Machblockgen.v Machblockgenproof.v\\
Asmblock.v Asmblockgen.v Asmblockgenproof0.v Asmblockgenproof1.v Asmblockgenproof.v Asmvliw.v\\
ForwardSimulationBlock.v PostpassScheduling.v PostpassSchedulingproof.v\\
- Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v ExtValues.v\\
+ Asmblockdeps.v DecBoolOps.v Chunks.v Peephole.v ExtValues.v ExtFloats.v\\
AbstractBasicBlocksDef.v SeqSimuTheory.v ImpSimuTest.v Parallelizability.v\\
ImpConfig.v ImpCore.v ImpExtern.v ImpHCons.v ImpIO.v ImpLoops.v ImpMonads.v ImpPrelude.v
EOF
diff --git a/cparser/Builtins.ml b/cparser/Builtins.ml
deleted file mode 100644
index 8eb1abfd..00000000
--- a/cparser/Builtins.ml
+++ /dev/null
@@ -1,54 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(* Compiler built-ins *)
-
-open C
-open Cutil
-
-let env = ref Env.empty
-let idents = ref []
-let decls = ref []
-
-let environment () = !env
-let identifiers () = !idents
-let declarations () = List.rev !decls
-
-let add_typedef (s, ty) =
- let (id, env') = Env.enter_typedef !env s ty in
- env := env';
- idents := id :: !idents;
- decls := {gdesc = Gtypedef(id, ty); gloc = no_loc} :: !decls
-
-let add_function (s, (res, args, va)) =
- let ty =
- TFun(res,
- Some (List.map (fun ty -> (Env.fresh_ident "", ty)) args),
- va, []) in
- let (id, env') = Env.enter_ident !env s Storage_extern ty in
- env := env';
- idents := id :: !idents;
- decls := {gdesc = Gdecl(Storage_extern, id, ty, None); gloc = no_loc} :: !decls
-
-type t = {
- typedefs: (string * C.typ) list;
- functions: (string * (C.typ * C.typ list * bool)) list
-}
-
-let set blt =
- env := Env.empty;
- idents := [];
- List.iter add_typedef blt.typedefs;
- List.iter add_function blt.functions
diff --git a/cparser/C.mli b/cparser/C.mli
index cc8d4065..15717565 100644
--- a/cparser/C.mli
+++ b/cparser/C.mli
@@ -264,3 +264,10 @@ and globdecl_desc =
| Gpragma of string (* #pragma directive *)
type program = globdecl list
+
+(** Builtin types and functions *)
+
+type builtins = {
+ builtin_typedefs: (string * typ) list;
+ builtin_functions: (string * (typ * typ list * bool)) list
+}
diff --git a/cparser/Cabs.v b/cparser/Cabs.v
index 5865ab69..5f12e8a1 100644
--- a/cparser/Cabs.v
+++ b/cparser/Cabs.v
@@ -20,7 +20,7 @@ Parameter string : Type.
(* OCaml's int64 type, used to represent individual characters in literals. *)
Parameter char_code : Type.
(* Context information. *)
-Parameter cabsloc : Type.
+Parameter loc : Type.
Record floatInfo := {
isHex_FI:bool;
@@ -51,7 +51,7 @@ Inductive typeSpecifier := (* Merge all specifiers into one type *)
* They also have a list of __attribute__s that appeared between the
* keyword and the type name (definitions only) *)
| Tstruct_union : structOrUnion -> option string -> option (list field_group) -> list attribute -> typeSpecifier
- | Tenum : option string -> option (list (string * option expression * cabsloc)) -> list attribute -> typeSpecifier
+ | Tenum : option string -> option (list (string * option expression * loc)) -> list attribute -> typeSpecifier
with storage :=
AUTO | STATIC | EXTERN | REGISTER | TYPEDEF
@@ -87,18 +87,18 @@ with decl_type :=
| PROTO_OLD : decl_type -> list string -> decl_type
with parameter :=
- | PARAM : list spec_elem -> option string -> decl_type -> list attribute -> cabsloc -> parameter
+ | PARAM : list spec_elem -> option string -> decl_type -> list attribute -> loc -> parameter
(* The optional expression is the bitfield *)
with field_group :=
- | Field_group : list spec_elem -> list (option name * option expression) -> cabsloc -> field_group
+ | Field_group : list spec_elem -> list (option name * option expression) -> loc -> field_group
(* The decl_type is in the order in which they are printed. Only the name of
* the declared identifier is pulled out. *)
(* e.g: in "int *x", "*x" is the declarator; "x" will be pulled out as *)
(* the string, and decl_type will be PTR([], JUSTBASE) *)
with name :=
- | Name : string -> decl_type -> list attribute -> cabsloc -> name
+ | Name : string -> decl_type -> list attribute -> loc -> name
(* A variable declarator ("name") with an initializer *)
with init_name :=
@@ -161,9 +161,9 @@ with initwhat :=
| ATINDEX_INIT : expression -> initwhat
with attribute :=
- | GCC_ATTR : list gcc_attribute -> cabsloc -> attribute
- | PACKED_ATTR : list expression -> cabsloc -> attribute
- | ALIGNAS_ATTR : list expression -> cabsloc -> attribute
+ | GCC_ATTR : list gcc_attribute -> loc -> attribute
+ | PACKED_ATTR : list expression -> loc -> attribute
+ | ALIGNAS_ATTR : list expression -> loc -> attribute
with gcc_attribute :=
| GCC_ATTR_EMPTY
@@ -194,31 +194,31 @@ Definition asm_flag := (bool * list char_code)%type.
** Declaration definition (at toplevel)
*)
Inductive definition :=
- | FUNDEF : list spec_elem -> name -> list definition -> statement -> cabsloc -> definition
- | DECDEF : init_name_group -> cabsloc -> definition (* global variable(s), or function prototype *)
- | PRAGMA : string -> cabsloc -> definition
+ | FUNDEF : list spec_elem -> name -> list definition -> statement -> loc -> definition
+ | DECDEF : init_name_group -> loc -> definition (* global variable(s), or function prototype *)
+ | PRAGMA : string -> loc -> definition
(*
** statements
*)
with statement :=
- | NOP : cabsloc -> statement
- | COMPUTATION : expression -> cabsloc -> statement
- | BLOCK : list statement -> cabsloc -> statement
- | If : expression -> statement -> option statement -> cabsloc -> statement
- | WHILE : expression -> statement -> cabsloc -> statement
- | DOWHILE : expression -> statement -> cabsloc -> statement
- | FOR : option for_clause -> option expression -> option expression -> statement -> cabsloc -> statement
- | BREAK : cabsloc -> statement
- | CONTINUE : cabsloc -> statement
- | RETURN : option expression -> cabsloc -> statement
- | SWITCH : expression -> statement -> cabsloc -> statement
- | CASE : expression -> statement -> cabsloc -> statement
- | DEFAULT : statement -> cabsloc -> statement
- | LABEL : string -> statement -> cabsloc -> statement
- | GOTO : string -> cabsloc -> statement
- | ASM : list cvspec -> bool -> list char_code -> list asm_operand -> list asm_operand -> list asm_flag -> cabsloc -> statement
+ | NOP : loc -> statement
+ | COMPUTATION : expression -> loc -> statement
+ | BLOCK : list statement -> loc -> statement
+ | If : expression -> statement -> option statement -> loc -> statement
+ | WHILE : expression -> statement -> loc -> statement
+ | DOWHILE : expression -> statement -> loc -> statement
+ | FOR : option for_clause -> option expression -> option expression -> statement -> loc -> statement
+ | BREAK : loc -> statement
+ | CONTINUE : loc -> statement
+ | RETURN : option expression -> loc -> statement
+ | SWITCH : expression -> statement -> loc -> statement
+ | CASE : expression -> statement -> loc -> statement
+ | DEFAULT : statement -> loc -> statement
+ | LABEL : string -> statement -> loc -> statement
+ | GOTO : string -> loc -> statement
+ | ASM : list cvspec -> bool -> list char_code -> list asm_operand -> list asm_operand -> list asm_flag -> loc -> statement
| DEFINITION : definition -> statement (*definition or declaration of a variable or type*)
with for_clause :=
diff --git a/cparser/Cabshelper.ml b/cparser/Cabshelper.ml
index 958f242c..22f3b3c7 100644
--- a/cparser/Cabshelper.ml
+++ b/cparser/Cabshelper.ml
@@ -16,11 +16,6 @@
open Cabs
-let cabslu = {lineno = -10;
- filename = "cabs loc unknown";
- byteno = -10;
- ident = 0}
-
(*********** HELPER FUNCTIONS **********)
let rec isStatic = function
@@ -44,13 +39,13 @@ let rec isTypedef = function
| _ :: rest -> isTypedef rest
-let get_definitionloc (d : definition) : cabsloc =
+let get_definitionloc (d : definition) : loc =
match d with
| FUNDEF(_, _, _, _, l) -> l
| DECDEF(_, l) -> l
| PRAGMA(_, l) -> l
-let get_statementloc (s : statement) : cabsloc =
+let get_statementloc (s : statement) : loc =
begin
match s with
| NOP(loc) -> loc
@@ -72,8 +67,8 @@ begin
| ASM(_,_,_,_,_,_,loc) -> loc
end
-let string_of_cabsloc l =
+let string_of_loc l =
Printf.sprintf "%s:%d" l.filename l.lineno
-let format_cabsloc pp l =
+let format_loc pp l =
Format.fprintf pp "%s:%d" l.filename l.lineno
diff --git a/cparser/Checks.ml b/cparser/Checks.ml
index a30cde7d..17caf19a 100644
--- a/cparser/Checks.ml
+++ b/cparser/Checks.ml
@@ -18,44 +18,68 @@ open Diagnostics
open Cutil
open Env
-let unknown_attrs loc attrs =
- let unknown attr =
- let attr_class = class_of_attribute attr in
- if attr_class = Attr_unknown then
- warning loc Unknown_attribute
- "unknown attribute '%s' ignored" (name_of_attribute attr) in
- List.iter unknown attrs
+(* AST traversal functions *)
-let unknown_attrs_typ env loc ty =
- let attr = attributes_of_type env ty in
- unknown_attrs loc attr
+let fold_over_stmt_loc ~(expr: 'a -> location -> exp -> 'a)
+ ~(decl: 'a -> location -> decl -> 'a)
+ (a: 'a) (s: stmt) : 'a =
+ let rec fold a s =
+ match s.sdesc with
+ | Sskip -> a
+ | Sbreak -> a
+ | Scontinue -> a
+ | Slabeled(_, s1) -> fold a s1
+ | Sgoto _ -> a
+ | Sreturn None -> a
+ | Sreturn (Some e) -> expr a s.sloc e
+ | Sasm(_, _, outs, ins, _) -> asm_operands (asm_operands a s.sloc outs) s.sloc ins
+ | Sdo e -> expr a s.sloc e
+ | Sif (e, s1, s2) -> fold (fold (expr a s.sloc e) s1) s2
+ | Sseq (s1, s2) -> fold (fold a s1) s2
+ | Sfor (s1, e, s2, s3) -> fold (fold (expr (fold a s1) s.sloc e) s2) s3
+ | Swhile(e, s1) -> fold (expr a s.sloc e) s1
+ | Sdowhile (s1, e) -> expr (fold a s1) s.sloc e
+ | Sswitch (e, s1) -> fold (expr a s.sloc e) s1
+ | Sblock sl -> List.fold_left fold a sl
+ | Sdecl d -> decl a s.sloc d
+ and asm_operands a loc l =
+ List.fold_left (fun a (_, _, e) -> expr a loc e) a l
+ in fold a s
-let unknown_attrs_decl env loc (sto, id, ty, init) =
- unknown_attrs_typ env loc ty
+let iter_over_stmt_loc
+ ?(expr = fun loc e -> ())
+ ?(decl = fun loc decl -> ())
+ (s: stmt) : unit =
+ fold_over_stmt_loc ~expr: (fun () loc e -> expr loc e)
+ ~decl: (fun () loc d -> decl loc d)
+ () s
+
+let fold_over_stmt ~(expr: 'a -> exp -> 'a)
+ ~(decl: 'a -> location -> decl -> 'a)
+ (a: 'a) (s: stmt) : 'a =
+ fold_over_stmt_loc ~expr:(fun a _ e -> expr a e) ~decl:decl a s
+
+let iter_over_stmt ?(expr = fun e -> ())
+ ?(decl = fun loc decl -> ())
+ (s:stmt) : unit =
+ fold_over_stmt_loc ~expr:(fun () _ e -> expr e)
+ ~decl:(fun () loc d -> decl loc d) () s
+
+let fold_over_init ~(expr: 'a -> exp -> 'a) (a: 'a) (i: init) : 'a =
+ let rec fold a = function
+ | Init_single e -> expr a e
+ | Init_array il -> List.fold_left fold a il
+ | Init_struct (_, sl) -> List.fold_left (fun a (_,i) -> fold a i) a sl
+ | Init_union (_, _, ui) -> fold a ui
+ in fold a i
-let rec unknown_attrs_stmt env s =
- match s.sdesc with
- | Sskip
- | Sbreak
- | Scontinue
- | Slabeled _
- | Sgoto _
- | Sreturn _
- | Sasm _
- | Sdo _ -> ()
- | Sif (_,s1,s2)
- | Sseq (s1,s2) ->
- unknown_attrs_stmt env s1;
- unknown_attrs_stmt env s2
- | Sfor (s1,e,s2,s3) ->
- unknown_attrs_stmt env s1;
- unknown_attrs_stmt env s2;
- unknown_attrs_stmt env s3
- | Swhile(_,s)
- | Sdowhile (s,_)
- | Sswitch (_,s) -> unknown_attrs_stmt env s
- | Sblock sl -> List.iter (unknown_attrs_stmt env) sl
- | Sdecl d -> unknown_attrs_decl env s.sloc d
+let iter_over_init ~(expr: exp -> unit) (i:init) : unit =
+ fold_over_init ~expr:(fun () e -> expr e) () i
+
+let fold_over_decl ~(expr: 'a -> exp -> 'a) (a: 'a) loc (sto, id, ty, init) : 'a=
+ match init with
+ | Some i -> fold_over_init ~expr a i
+ | None -> a
let traverse_program
?(decl = fun env loc d -> ())
@@ -93,7 +117,27 @@ let traverse_program
pragma env g.gloc s;
env in
traverse env gl in
- traverse (Builtins.environment ()) p
+ traverse (Env.initial ()) p
+
+(* Unknown attributes warning *)
+
+let unknown_attrs loc attrs =
+ let unknown attr =
+ let attr_class = class_of_attribute attr in
+ if attr_class = Attr_unknown then
+ warning loc Unknown_attribute
+ "unknown attribute '%s' ignored" (name_of_attribute attr) in
+ List.iter unknown attrs
+
+let unknown_attrs_typ env loc ty =
+ let attr = attributes_of_type env ty in
+ unknown_attrs loc attr
+
+let unknown_attrs_decl env loc (sto, id, ty, init) =
+ unknown_attrs_typ env loc ty
+
+let unknown_attrs_stmt env s =
+ iter_over_stmt ~decl:(unknown_attrs_decl env) s
let unknown_attrs_program p =
let decl env loc d =
@@ -122,6 +166,7 @@ let unknown_attrs_program p =
~enum:enum
p
+(* Unused variables and parameters warning *)
let rec vars_used_expr env e =
match e.edesc with
@@ -143,83 +188,21 @@ let rec vars_used_expr env e =
let env = vars_used_expr env e in
List.fold_left vars_used_expr env p
-and vars_used_init env = function
- | Init_single e -> vars_used_expr env e
- | Init_array al -> List.fold_left vars_used_init env al
- | Init_struct (_,sl) -> List.fold_left (fun env (_,i) -> vars_used_init env i) env sl
- | Init_union (_,_,ui) -> vars_used_init env ui
-
-let rec vars_used_stmt env s =
- match s.sdesc with
- | Sbreak
- | Scontinue
- | Sgoto _
- | Sreturn None
- | Sskip -> env
- | Sreturn (Some e)
- | Sdo e -> (vars_used_expr env e)
- | Sseq (s1,s2) ->
- let env = vars_used_stmt env s1 in
- vars_used_stmt env s2
- | Sif (e,s1,s2) ->
- let env = vars_used_expr env e in
- let env = vars_used_stmt env s1 in
- vars_used_stmt env s2
- | Sfor (s1,e,s2,s3) ->
- let env = vars_used_expr env e in
- let env = vars_used_stmt env s1 in
- let env = vars_used_stmt env s2 in
- vars_used_stmt env s3
- | Sswitch (e,s)
- | Swhile (e,s)
- | Sdowhile (s,e) ->
- let env = vars_used_expr env e in
- vars_used_stmt env s
- | Sblock sl -> List.fold_left vars_used_stmt env sl
- | Sdecl (sto,id,ty,init) ->
- let env = match init with
- | Some init ->vars_used_init env init
- | None -> env in
- env
- | Slabeled (lbl,s) -> vars_used_stmt env s
- | Sasm (attr,str,op,op2,constr) ->
- let vars_asm_op env (_,_,e) =
- vars_used_expr env e in
- let env = List.fold_left vars_asm_op env op in
- let env = List.fold_left vars_asm_op env op2 in
- env
-
-let unused_variable env used loc (id,ty) =
+and vars_used_init env init =
+ fold_over_init ~expr:vars_used_expr env init
+
+let vars_used_stmt env s =
+ fold_over_stmt ~expr: vars_used_expr
+ ~decl: (fold_over_decl ~expr: vars_used_expr) env s
+
+let unused_variable env used loc (id, ty) =
let attr = attributes_of_type env ty in
let unused_attr = find_custom_attributes ["unused";"__unused__"] attr <> [] in
if not ((IdentSet.mem id used) || unused_attr) then
warning loc Unused_variable "unused variable '%s'" id.name
-let rec unused_variables_stmt env used s =
- match s.sdesc with
- | Sbreak
- | Scontinue
- | Sgoto _
- | Sreturn _
- | Sskip
- | Sasm _
- | Sdo _ -> ()
- | Sseq (s1,s2)
- | Sif (_,s1,s2) ->
- unused_variables_stmt env used s1;
- unused_variables_stmt env used s2
- | Sfor (s1,e,s2,s3) ->
- unused_variables_stmt env used s1;
- unused_variables_stmt env used s2;
- unused_variables_stmt env used s3
- | Slabeled (_,s)
- | Sswitch (_,s)
- | Swhile (_,s)
- | Sdowhile (s,_) ->
- unused_variables_stmt env used s
- | Sblock sl -> List.iter (unused_variables_stmt env used) sl
- | Sdecl (sto,id,ty,init) ->
- unused_variable env used s.sloc (id,ty)
+let unused_variables_stmt env used s =
+ iter_over_stmt ~decl:(fun loc (sto, id, ty, init) -> unused_variable env used loc (id,ty)) s
let unused_variables p =
let fundef env loc fd =
@@ -229,3 +212,166 @@ let unused_variables p =
traverse_program
~fundef:fundef
p
+
+(* Warning for conditionals that cannot be transformed into linear code *)
+
+(* Compute the set of local variables that do not have their address taken *)
+
+let rec non_stack_locals_expr vars e =
+ match e.edesc with
+ | ECast (_,e) -> non_stack_locals_expr vars e
+ | EUnop (Oaddrof,e) ->
+ begin match e.edesc with
+ | EVar id ->
+ IdentSet.remove id vars
+ | _ -> vars
+ end
+ | EUnop (Oderef, e) ->
+ (* Special optimization *(& ...) is removed in SimplExpr *)
+ begin match e.edesc with
+ | EUnop (Oaddrof,e) -> non_stack_locals_expr vars e
+ | _ -> non_stack_locals_expr vars e
+ end
+ | EUnop (_, e) ->
+ non_stack_locals_expr vars e
+ | EBinop (_,e1,e2,_) ->
+ let vars = non_stack_locals_expr vars e1 in
+ non_stack_locals_expr vars e2
+ | EConditional (e1,e2,e3) ->
+ let vars = non_stack_locals_expr vars e1 in
+ let vars = non_stack_locals_expr vars e2 in
+ non_stack_locals_expr vars e3
+ | ECompound (_,init) -> non_stack_locals_init vars init
+ | ECall (e,p) ->
+ let vars = non_stack_locals_expr vars e in
+ List.fold_left non_stack_locals_expr vars p
+ | _ -> vars
+
+and non_stack_locals_init vars init =
+ fold_over_init ~expr:non_stack_locals_expr vars init
+
+let add_vars env vars (id,ty) =
+ let volatile = List.mem AVolatile (attributes_of_type env ty) in
+ if not volatile then
+ IdentSet.add id vars
+ else
+ vars
+
+let non_stack_locals_stmt env vars s =
+ let decl vars loc (sto, id, ty, init) =
+ let vars = match init with
+ | Some init -> non_stack_locals_init vars init
+ | None -> vars in
+ add_vars env vars (id,ty) in
+ fold_over_stmt ~expr:non_stack_locals_expr ~decl:decl
+ vars s
+
+(* Check whether an expression is safe and can be always evaluated *)
+
+let safe_cast env tfrom tto =
+ match unroll env tfrom, unroll env tto with
+ | (TInt _ | TPtr _ | TArray _ | TFun _ | TEnum _),
+ (TInt _ | TPtr _ | TEnum _) -> true
+ | TFloat _, TFloat _ -> true
+ | _, _ -> equal_types env tfrom tto
+
+let safe_expr vars env e =
+ let rec expr e =
+ match e.edesc with
+ | EConst _ | ESizeof _ | EAlignof _ | ECompound _ -> true
+ | EVar id -> (IdentSet.mem id vars) || not (is_scalar_type env e.etyp)
+ | ECast (ty, e) ->
+ safe_cast env e.etyp ty && expr e
+ | EUnop (op, e) ->
+ unop op e
+ | EBinop (op, e1, e2, ty) ->
+ binop op e1 e2
+ | EConditional _ -> false
+ | ECall _ -> false
+ and binop op e1 e2 =
+ let is_long_long_type ty =
+ match unroll env ty with
+ | TInt (ILongLong, _)
+ | TInt (IULongLong, _) -> true
+ | _ -> false in
+ match op with
+ | Oadd | Osub | Omul | Oand | Oor | Oxor | Oshl | Oshr ->
+ expr e1 && expr e2
+ | Oeq | One | Olt | Ogt | Ole | Oge ->
+ let not_long_long = not (is_long_long_type e1.etyp) && not (is_long_long_type e2.etyp) in
+ not_long_long && expr e1 && expr e2
+ | _ -> false
+ (* x.f if f has array or struct or union type *)
+ and unop op e =
+ match op with
+ | Ominus | Onot | Olognot | Oplus -> expr e
+ | Oaddrof ->
+ begin match e.edesc with
+ (* skip &*e *)
+ | EUnop (Oderef, e) -> expr e
+ (* skip &(e.f) *)
+ | EUnop (Odot f, e) -> expr e
+ | _ -> expr e
+ end
+ (* skip *&e *)
+ | Oderef ->
+ begin match e.edesc with
+ | EUnop (Oaddrof,e) -> expr e
+ | _ -> false
+ end
+ (* e.f is okay if f has array or composite type *)
+ | Odot m ->
+ let fld = field_of_dot_access env e.etyp m in
+ (is_array_type env fld.fld_typ || is_composite_type env fld.fld_typ) && expr e
+ | _ -> false in
+ expr e
+
+(* Check expressions if they contain conditionals that cannot be transformed in
+ linear code. The inner_cond parameter is used to mimic the translation of short
+ circuit logical or and logical and as well as conditional to if statements in
+ SimplExpr. *)
+
+let rec non_linear_cond_expr inner_cond vars env loc e =
+ match e.edesc with
+ | EConst _ | ESizeof _ | EAlignof _ | EVar _ -> ()
+ | ECast (_ , e) | EUnop (_, e)-> non_linear_cond_expr false vars env loc e
+ | EBinop (op, e1, e2, ty) ->
+ let inner_cond = match op with
+ | Ocomma -> inner_cond
+ | Ologand | Ologor -> true
+ | _ -> false
+ in
+ non_linear_cond_expr false vars env loc e1;
+ non_linear_cond_expr inner_cond vars env loc e2
+ | EConditional (c, e1, e2) ->
+ let can_cast = safe_cast env e1.etyp e.etyp && safe_cast env e2.etyp e.etyp in
+ if not can_cast || inner_cond || not (safe_expr vars env e1) || not (safe_expr vars env e2) then
+ warning loc Non_linear_cond_expr "conditional expression may not be linearized";
+ non_linear_cond_expr true vars env loc e1;
+ non_linear_cond_expr true vars env loc e2;
+ | ECompound (ty, init) -> non_linear_cond_init vars env loc init
+ | ECall (e, params) ->
+ non_linear_cond_expr false vars env loc e;
+ List.iter (non_linear_cond_expr false vars env loc) params
+
+and non_linear_cond_init vars env loc init =
+ iter_over_init ~expr:(non_linear_cond_expr false vars env loc) init
+
+let non_linear_cond_stmt vars env s =
+ let decl loc (sto, id, ty, init) =
+ match init with
+ | None -> ()
+ | Some init -> non_linear_cond_init vars env loc init in
+ iter_over_stmt_loc ~expr:(non_linear_cond_expr false vars env) ~decl:decl s
+
+let non_linear_conditional p =
+ if active_warning Non_linear_cond_expr && !Clflags.option_Obranchless then begin
+ let fundef env loc fd =
+ let vars = List.fold_left (add_vars env) IdentSet.empty fd.fd_params in
+ let vars = non_stack_locals_stmt env vars fd.fd_body in
+ non_linear_cond_stmt vars env fd.fd_body;
+ in
+ traverse_program
+ ~fundef:fundef
+ p
+ end
diff --git a/cparser/Checks.mli b/cparser/Checks.mli
index 4d61a5b8..cfd7b04d 100644
--- a/cparser/Checks.mli
+++ b/cparser/Checks.mli
@@ -16,3 +16,5 @@
val unknown_attrs_program: C.program -> unit
val unused_variables: C.program -> unit
+
+val non_linear_conditional : C.program -> unit
diff --git a/cparser/Cutil.ml b/cparser/Cutil.ml
index 7329767a..7a2f4828 100644
--- a/cparser/Cutil.ml
+++ b/cparser/Cutil.ml
@@ -29,7 +29,7 @@ let no_loc = ("", -1)
module Ident = struct
type t = ident
- let compare id1 id2 = Pervasives.compare id1.stamp id2.stamp
+ let compare id1 id2 = compare id1.stamp id2.stamp
end
module IdentSet = Set.Make(Ident)
@@ -821,6 +821,11 @@ let is_composite_type env t =
| TStruct _ | TUnion _ -> true
| _ -> false
+let is_array_type env t =
+ match unroll env t with
+ | TArray _ -> true
+ | _ -> false
+
let is_function_type env t =
match unroll env t with
| TFun _ -> true
diff --git a/cparser/Cutil.mli b/cparser/Cutil.mli
index 5a1e9af3..f6c4627d 100644
--- a/cparser/Cutil.mli
+++ b/cparser/Cutil.mli
@@ -166,6 +166,8 @@ val is_scalar_type : Env.t -> typ -> bool
(* Is type integer, float or pointer? *)
val is_composite_type : Env.t -> typ -> bool
(* Is type a struct or union? *)
+val is_array_type : Env.t -> typ -> bool
+ (* Is type an array type? *)
val is_function_type : Env.t -> typ -> bool
(* Is type a function type? (not pointer to function) *)
val is_function_pointer_type : Env.t -> typ -> bool
diff --git a/cparser/Diagnostics.ml b/cparser/Diagnostics.ml
index 51dcab47..012e4b66 100644
--- a/cparser/Diagnostics.ml
+++ b/cparser/Diagnostics.ml
@@ -102,6 +102,7 @@ type warning_type =
| Flexible_array_extensions
| Tentative_incomplete_static
| Reduced_alignment
+ | Non_linear_cond_expr
(* List of active warnings *)
let active_warnings: warning_type list ref = ref [
@@ -163,6 +164,7 @@ let string_of_warning = function
| Flexible_array_extensions -> "flexible-array-extensions"
| Tentative_incomplete_static -> "tentative-incomplete-static"
| Reduced_alignment -> "reduced-alignment"
+ | Non_linear_cond_expr -> "non-linear-cond-expr"
(* Activate the given warning *)
let activate_warning w () =
@@ -216,6 +218,7 @@ let wall () =
Flexible_array_extensions;
Tentative_incomplete_static;
Reduced_alignment;
+ Non_linear_cond_expr;
]
let wnothing () =
@@ -253,6 +256,7 @@ let werror () =
Flexible_array_extensions;
Tentative_incomplete_static;
Reduced_alignment;
+ Non_linear_cond_expr;
]
(* Generate the warning key for the message *)
@@ -437,6 +441,7 @@ let warning_options =
error_option Flexible_array_extensions @
error_option Tentative_incomplete_static @
error_option Reduced_alignment @
+ error_option Non_linear_cond_expr @
[Exact ("-Wfatal-errors"), Set error_fatal;
Exact ("-fdiagnostics-color"), Ignore; (* Either output supports it or no color *)
Exact ("-fno-diagnostics-color"), Unset color_diagnostics;
@@ -492,3 +497,6 @@ let crash exn =
let no_loc = ("", -1)
let file_loc file = (file,-10)
+
+let active_warning ty =
+ fst (classify_warning ty) <> SuppressedMsg
diff --git a/cparser/Diagnostics.mli b/cparser/Diagnostics.mli
index 6a3c11c8..0f0a0ea5 100644
--- a/cparser/Diagnostics.mli
+++ b/cparser/Diagnostics.mli
@@ -55,6 +55,7 @@ type warning_type =
| Flexible_array_extensions (** usange of structs with flexible arrays in structs and arrays *)
| Tentative_incomplete_static (** static tentative definition with incomplete type *)
| Reduced_alignment (** alignment reduction *)
+ | Non_linear_cond_expr (** condition that cannot be linearized *)
val warning : (string * int) -> warning_type -> ('a, Format.formatter, unit, unit, unit, unit) format6 -> 'a
(** [warning (f,c) w fmt arg1 ... argN] formats the arguments [arg1] to [argN] as warining according to
@@ -95,3 +96,6 @@ val file_loc : string -> string * int
val error_summary : unit -> unit
(** Print a summary containing the numbers of errors encountered *)
+
+val active_warning : warning_type -> bool
+(** Test whether a warning is active to avoid costly checks *)
diff --git a/cparser/Elab.ml b/cparser/Elab.ml
index 10380152..3797164d 100644
--- a/cparser/Elab.ml
+++ b/cparser/Elab.ml
@@ -258,7 +258,7 @@ let enter_or_refine_function loc env id sto ty =
(* Forward declarations *)
-let elab_expr_f : (cabsloc -> Env.t -> Cabs.expression -> C.exp * Env.t) ref
+let elab_expr_f : (Cabs.loc -> Env.t -> Cabs.expression -> C.exp * Env.t) ref
= ref (fun _ _ _ -> assert false)
let elab_funbody_f : (C.typ -> bool -> bool -> Env.t -> statement -> C.stmt) ref
@@ -1802,6 +1802,42 @@ let elab_expr ctx loc env a =
(print_typ env) ty (print_typ env) ty' (print_typ env) ty' (print_typ env) ty;
{ edesc = ECall(ident, [b2; b3]); etyp = ty },env
+ | CALL((VARIABLE "__builtin_sel" as a0), al) ->
+ begin match al with
+ | [a1; a2; a3] ->
+ let b0,env = elab env a0 in
+ let b1,env = elab env a1 in
+ let b2,env = elab env a2 in
+ let b3,env = elab env a3 in
+ if not (is_scalar_type env b1.etyp) then
+ error "first argument of '__builtin_sel' is not a scalar type (invalid %a)"
+ (print_typ env) b1.etyp;
+ let tyres =
+ match pointer_decay env b2.etyp, pointer_decay env b3.etyp with
+ | (TInt _ | TFloat _ | TEnum _), (TInt _ | TFloat _ | TEnum _) ->
+ binary_conversion env b2.etyp b3.etyp
+ | (TPtr(ty1, a1) as pty1), (TPtr(ty2, a2) as pty2) ->
+ if is_void_type env ty1 || is_void_type env ty2 then
+ TPtr(TVoid (add_attributes a1 a2), [])
+ else begin
+ match combine_types AttrIgnoreAll env pty1 pty2 with
+ | None ->
+ warning Pointer_type_mismatch "the second and third arguments of '__builtin_sel' have incompatible pointer types (%a and %a)"
+ (print_typ env) pty1 (print_typ env) pty2;
+ (* tolerance *)
+ TPtr(TVoid (add_attributes a1 a2), [])
+ | Some ty -> ty
+ end
+ | _, _ ->
+ fatal_error "wrong types (%a and %a) for the second and third arguments of '__builtin_sel'"
+ (print_typ env) b2.etyp (print_typ env) b3.etyp
+
+ in
+ { edesc = ECall(b0, [b1; b2; b3]); etyp = tyres }, env
+ | _ ->
+ fatal_error "'__builtin_sel' expect 3 arguments"
+ end
+
| CALL(a1, al) ->
let b1,env =
(* Catch the old-style usage of calling a function without
@@ -2708,7 +2744,7 @@ let elab_fundef genv spec name defs body loc =
(* Definitions *)
let elab_decdef (for_loop: bool) (local: bool) (nonstatic_inline: bool)
(env: Env.t) ((spec, namelist): Cabs.init_name_group)
- (loc: Cabs.cabsloc) : decl list * Env.t =
+ (loc: Cabs.loc) : decl list * Env.t =
let (sto, inl, noret, tydef, bty, env') =
elab_specifier ~only:(namelist=[]) loc env spec in
(* Sanity checks on storage class *)
@@ -3089,10 +3125,11 @@ let _ = elab_funbody_f := elab_funbody
let elab_file prog =
reset();
- let env = Builtins.environment () in
+ let env = Env.initial () in
let elab_def env d = snd (elab_definition false false false env d) in
ignore (List.fold_left elab_def env prog);
let p = elaborated_program () in
Checks.unused_variables p;
Checks.unknown_attrs_program p;
+ Checks.non_linear_conditional p;
p
diff --git a/cparser/Elab.mli b/cparser/Elab.mli
index f701e8c5..59c5efc1 100644
--- a/cparser/Elab.mli
+++ b/cparser/Elab.mli
@@ -18,8 +18,8 @@ val elab_file : Cabs.definition list -> C.program
definitions as produced by the parser into a program in C abstract
syntax. *)
-val elab_int_constant : Cabs.cabsloc -> string -> int64 * C.ikind
+val elab_int_constant : Cabs.loc -> string -> int64 * C.ikind
val elab_float_constant : Cabs.floatInfo -> C.float_cst * C.fkind
-val elab_char_constant : Cabs.cabsloc -> bool -> int64 list -> int64
+val elab_char_constant : Cabs.loc -> bool -> int64 list -> int64
(* These auxiliary functions are exported so that they can be reused
in other projects that deal with C-style source languages. *)
diff --git a/cparser/Env.ml b/cparser/Env.ml
index 5fa4571a..4723a725 100644
--- a/cparser/Env.ml
+++ b/cparser/Env.ml
@@ -276,6 +276,46 @@ let add_enum env id info =
let add_types env_old env_new =
{ env_new with env_ident = env_old.env_ident;env_scope = env_old.env_scope;}
+(* Initial environment describing the built-in types and functions *)
+
+module Init = struct
+
+let env = ref empty
+let idents = ref []
+let decls = ref []
+
+let no_loc = ("", -1)
+
+let add_typedef (s, ty) =
+ let (id, env') = enter_typedef !env s ty in
+ env := env';
+ idents := id :: !idents;
+ decls := {gdesc = Gtypedef(id, ty); gloc = no_loc} :: !decls
+
+let add_function (s, (res, args, va)) =
+ let ty =
+ TFun(res,
+ Some (List.map (fun ty -> (fresh_ident "", ty)) args),
+ va, []) in
+ let (id, env') = enter_ident !env s Storage_extern ty in
+ env := env';
+ idents := id :: !idents;
+ decls :=
+ {gdesc = Gdecl(Storage_extern, id, ty, None); gloc = no_loc} :: !decls
+
+end
+
+let initial () = !Init.env
+let initial_identifiers () = !Init.idents
+let initial_declarations () = List.rev !Init.decls
+
+let set_builtins blt =
+ Init.env := empty;
+ Init.idents := [];
+ Init.decls := [];
+ List.iter Init.add_typedef blt.builtin_typedefs;
+ List.iter Init.add_function blt.builtin_functions
+
(* Error reporting *)
open Printf
diff --git a/cparser/Env.mli b/cparser/Env.mli
index 7ea2c514..1baab68f 100644
--- a/cparser/Env.mli
+++ b/cparser/Env.mli
@@ -77,3 +77,10 @@ val add_typedef : t -> C.ident -> typedef_info -> t
val add_enum : t -> C.ident -> enum_info -> t
val add_types : t -> t -> t
+
+(* Initial environment describing the builtin types and functions *)
+
+val initial: unit -> t
+val initial_identifiers: unit -> C.ident list
+val initial_declarations: unit -> C.globdecl list
+val set_builtins: C.builtins -> unit
diff --git a/cparser/GCC.ml b/cparser/GCC.ml
index 010d12f3..458e51d3 100644
--- a/cparser/GCC.ml
+++ b/cparser/GCC.ml
@@ -38,10 +38,10 @@ let intPtrType = TPtr(TInt(IInt, []), [])
let sizeType() = TInt(size_t_ikind(), [])
let builtins = {
- Builtins.typedefs = [
+ builtin_typedefs = [
"__builtin_va_list", voidPtrType
];
- Builtins.functions = [
+ builtin_functions = [
"__builtin___fprintf_chk", (intType, [ voidPtrType; intType; charConstPtrType ], true) (* first argument is really FILE*, not void*, but we don't want to build in the definition for FILE *);
"__builtin___memcpy_chk", (voidPtrType, [ voidPtrType; voidConstPtrType; sizeType(); sizeType() ], false);
"__builtin___memmove_chk", (voidPtrType, [ voidPtrType; voidConstPtrType; sizeType(); sizeType() ], false);
diff --git a/cparser/Lexer.mll b/cparser/Lexer.mll
index 7cf22eef..346477b5 100644
--- a/cparser/Lexer.mll
+++ b/cparser/Lexer.mll
@@ -20,7 +20,7 @@ open Pre_parser_aux
module SSet = Set.Make(String)
-let lexicon : (string, Cabs.cabsloc -> token) Hashtbl.t = Hashtbl.create 17
+let lexicon : (string, Cabs.loc -> token) Hashtbl.t = Hashtbl.create 17
let ignored_keywords : SSet.t ref = ref SSet.empty
let reserved_keyword loc id =
@@ -434,10 +434,7 @@ and singleline_comment = parse
| _ { singleline_comment lexbuf }
{
- open Streams
- open Specif
- open Parser
- open !Aut.GramDefs
+ open Parser.MenhirLibParser.Inter
(* This is the main entry point to the lexer. *)
@@ -463,8 +460,8 @@ and singleline_comment = parse
curr_id := None;
let loc = currentLoc lexbuf in
let token =
- if SSet.mem id !types_context then TYPEDEF_NAME (id, ref TypedefId, loc)
- else VAR_NAME (id, ref VarId, loc)
+ if SSet.mem id !types_context then Pre_parser.TYPEDEF_NAME (id, ref TypedefId, loc)
+ else Pre_parser.VAR_NAME (id, ref VarId, loc)
in
Queue.push token tokens;
token
@@ -497,133 +494,129 @@ and singleline_comment = parse
(* [tokens_stream filename text] runs the pre_parser and produces a stream
of (appropriately classified) tokens. *)
- let tokens_stream filename text : token coq_Stream =
+ let tokens_stream filename text : buffer =
let tokens = Queue.create () in
let buffer = ref ErrorReports.Zero in
invoke_pre_parser filename text (lexer tokens buffer) buffer;
- let rec compute_token_stream () =
- let loop t v =
- Cons (Coq_existT (t, Obj.magic v), Lazy.from_fun compute_token_stream)
- in
+ let rec compute_buffer () =
+ let loop t = Buf_cons (t, Lazy.from_fun compute_buffer) in
match Queue.pop tokens with
- | ADD_ASSIGN loc -> loop ADD_ASSIGN't loc
- | AND loc -> loop AND't loc
- | ANDAND loc -> loop ANDAND't loc
- | AND_ASSIGN loc -> loop AND_ASSIGN't loc
- | AUTO loc -> loop AUTO't loc
- | BANG loc -> loop BANG't loc
- | BAR loc -> loop BAR't loc
- | BARBAR loc -> loop BARBAR't loc
- | UNDERSCORE_BOOL loc -> loop UNDERSCORE_BOOL't loc
- | BREAK loc -> loop BREAK't loc
- | BUILTIN_VA_ARG loc -> loop BUILTIN_VA_ARG't loc
- | BUILTIN_OFFSETOF loc -> loop BUILTIN_OFFSETOF't loc
- | CASE loc -> loop CASE't loc
- | CHAR loc -> loop CHAR't loc
- | COLON loc -> loop COLON't loc
- | COMMA loc -> loop COMMA't loc
- | CONST loc -> loop CONST't loc
- | CONSTANT (cst, loc) -> loop CONSTANT't (cst, loc)
- | CONTINUE loc -> loop CONTINUE't loc
- | DEC loc -> loop DEC't loc
- | DEFAULT loc -> loop DEFAULT't loc
- | DIV_ASSIGN loc -> loop DIV_ASSIGN't loc
- | DO loc -> loop DO't loc
- | DOT loc -> loop DOT't loc
- | DOUBLE loc -> loop DOUBLE't loc
- | ELLIPSIS loc -> loop ELLIPSIS't loc
- | ELSE loc -> loop ELSE't loc
- | ENUM loc -> loop ENUM't loc
- | EOF -> loop EOF't ()
- | EQ loc -> loop EQ't loc
- | EQEQ loc -> loop EQEQ't loc
- | EXTERN loc -> loop EXTERN't loc
- | FLOAT loc -> loop FLOAT't loc
- | FOR loc -> loop FOR't loc
- | GEQ loc -> loop GEQ't loc
- | GOTO loc -> loop GOTO't loc
- | GT loc -> loop GT't loc
- | HAT loc -> loop HAT't loc
- | IF loc -> loop IF't loc
- | INC loc -> loop INC't loc
- | INLINE loc -> loop INLINE't loc
- | INT loc -> loop INT't loc
- | LBRACE loc -> loop LBRACE't loc
- | LBRACK loc -> loop LBRACK't loc
- | LEFT loc -> loop LEFT't loc
- | LEFT_ASSIGN loc -> loop LEFT_ASSIGN't loc
- | LEQ loc -> loop LEQ't loc
- | LONG loc -> loop LONG't loc
- | LPAREN loc -> loop LPAREN't loc
- | LT loc -> loop LT't loc
- | MINUS loc -> loop MINUS't loc
- | MOD_ASSIGN loc -> loop MOD_ASSIGN't loc
- | MUL_ASSIGN loc -> loop MUL_ASSIGN't loc
- | NEQ loc -> loop NEQ't loc
- | NORETURN loc -> loop NORETURN't loc
- | OR_ASSIGN loc -> loop OR_ASSIGN't loc
- | PACKED loc -> loop PACKED't loc
- | PERCENT loc -> loop PERCENT't loc
- | PLUS loc -> loop PLUS't loc
- | PTR loc -> loop PTR't loc
- | QUESTION loc -> loop QUESTION't loc
- | RBRACE loc -> loop RBRACE't loc
- | RBRACK loc -> loop RBRACK't loc
- | REGISTER loc -> loop REGISTER't loc
- | RESTRICT loc -> loop RESTRICT't loc
- | RETURN loc -> loop RETURN't loc
- | RIGHT loc -> loop RIGHT't loc
- | RIGHT_ASSIGN loc -> loop RIGHT_ASSIGN't loc
- | RPAREN loc -> loop RPAREN't loc
- | SEMICOLON loc -> loop SEMICOLON't loc
- | SHORT loc -> loop SHORT't loc
- | SIGNED loc -> loop SIGNED't loc
- | SIZEOF loc -> loop SIZEOF't loc
- | SLASH loc -> loop SLASH't loc
- | STAR loc -> loop STAR't loc
- | STATIC loc -> loop STATIC't loc
- | STRING_LITERAL (wide, str, loc) ->
+ | Pre_parser.ADD_ASSIGN loc -> loop (Parser.ADD_ASSIGN loc)
+ | Pre_parser.AND loc -> loop (Parser.AND loc)
+ | Pre_parser.ANDAND loc -> loop (Parser.ANDAND loc)
+ | Pre_parser.AND_ASSIGN loc -> loop (Parser.AND_ASSIGN loc)
+ | Pre_parser.AUTO loc -> loop (Parser.AUTO loc)
+ | Pre_parser.BANG loc -> loop (Parser.BANG loc)
+ | Pre_parser.BAR loc -> loop (Parser.BAR loc)
+ | Pre_parser.BARBAR loc -> loop (Parser.BARBAR loc)
+ | Pre_parser.UNDERSCORE_BOOL loc -> loop (Parser.UNDERSCORE_BOOL loc)
+ | Pre_parser.BREAK loc -> loop (Parser.BREAK loc)
+ | Pre_parser.BUILTIN_VA_ARG loc -> loop (Parser.BUILTIN_VA_ARG loc)
+ | Pre_parser.BUILTIN_OFFSETOF loc -> loop (Parser.BUILTIN_OFFSETOF loc)
+ | Pre_parser.CASE loc -> loop (Parser.CASE loc)
+ | Pre_parser.CHAR loc -> loop (Parser.CHAR loc)
+ | Pre_parser.COLON loc -> loop (Parser.COLON loc)
+ | Pre_parser.COMMA loc -> loop (Parser.COMMA loc)
+ | Pre_parser.CONST loc -> loop (Parser.CONST loc)
+ | Pre_parser.CONSTANT (cst, loc) -> loop (Parser.CONSTANT (cst, loc))
+ | Pre_parser.CONTINUE loc -> loop (Parser.CONTINUE loc)
+ | Pre_parser.DEC loc -> loop (Parser.DEC loc)
+ | Pre_parser.DEFAULT loc -> loop (Parser.DEFAULT loc)
+ | Pre_parser.DIV_ASSIGN loc -> loop (Parser.DIV_ASSIGN loc)
+ | Pre_parser.DO loc -> loop (Parser.DO loc)
+ | Pre_parser.DOT loc -> loop (Parser.DOT loc)
+ | Pre_parser.DOUBLE loc -> loop (Parser.DOUBLE loc)
+ | Pre_parser.ELLIPSIS loc -> loop (Parser.ELLIPSIS loc)
+ | Pre_parser.ELSE loc -> loop (Parser.ELSE loc)
+ | Pre_parser.ENUM loc -> loop (Parser.ENUM loc)
+ | Pre_parser.EOF -> loop (Parser.EOF ())
+ | Pre_parser.EQ loc -> loop (Parser.EQ loc)
+ | Pre_parser.EQEQ loc -> loop (Parser.EQEQ loc)
+ | Pre_parser.EXTERN loc -> loop (Parser.EXTERN loc)
+ | Pre_parser.FLOAT loc -> loop (Parser.FLOAT loc)
+ | Pre_parser.FOR loc -> loop (Parser.FOR loc)
+ | Pre_parser.GEQ loc -> loop (Parser.GEQ loc)
+ | Pre_parser.GOTO loc -> loop (Parser.GOTO loc)
+ | Pre_parser.GT loc -> loop (Parser.GT loc)
+ | Pre_parser.HAT loc -> loop (Parser.HAT loc)
+ | Pre_parser.IF loc -> loop (Parser.IF_ loc)
+ | Pre_parser.INC loc -> loop (Parser.INC loc)
+ | Pre_parser.INLINE loc -> loop (Parser.INLINE loc)
+ | Pre_parser.INT loc -> loop (Parser.INT loc)
+ | Pre_parser.LBRACE loc -> loop (Parser.LBRACE loc)
+ | Pre_parser.LBRACK loc -> loop (Parser.LBRACK loc)
+ | Pre_parser.LEFT loc -> loop (Parser.LEFT loc)
+ | Pre_parser.LEFT_ASSIGN loc -> loop (Parser.LEFT_ASSIGN loc)
+ | Pre_parser.LEQ loc -> loop (Parser.LEQ loc)
+ | Pre_parser.LONG loc -> loop (Parser.LONG loc)
+ | Pre_parser.LPAREN loc -> loop (Parser.LPAREN loc)
+ | Pre_parser.LT loc -> loop (Parser.LT loc)
+ | Pre_parser.MINUS loc -> loop (Parser.MINUS loc)
+ | Pre_parser.MOD_ASSIGN loc -> loop (Parser.MOD_ASSIGN loc)
+ | Pre_parser.MUL_ASSIGN loc -> loop (Parser.MUL_ASSIGN loc)
+ | Pre_parser.NEQ loc -> loop (Parser.NEQ loc)
+ | Pre_parser.NORETURN loc -> loop (Parser.NORETURN loc)
+ | Pre_parser.OR_ASSIGN loc -> loop (Parser.OR_ASSIGN loc)
+ | Pre_parser.PACKED loc -> loop (Parser.PACKED loc)
+ | Pre_parser.PERCENT loc -> loop (Parser.PERCENT loc)
+ | Pre_parser.PLUS loc -> loop (Parser.PLUS loc)
+ | Pre_parser.PTR loc -> loop (Parser.PTR loc)
+ | Pre_parser.QUESTION loc -> loop (Parser.QUESTION loc)
+ | Pre_parser.RBRACE loc -> loop (Parser.RBRACE loc)
+ | Pre_parser.RBRACK loc -> loop (Parser.RBRACK loc)
+ | Pre_parser.REGISTER loc -> loop (Parser.REGISTER loc)
+ | Pre_parser.RESTRICT loc -> loop (Parser.RESTRICT loc)
+ | Pre_parser.RETURN loc -> loop (Parser.RETURN loc)
+ | Pre_parser.RIGHT loc -> loop (Parser.RIGHT loc)
+ | Pre_parser.RIGHT_ASSIGN loc -> loop (Parser.RIGHT_ASSIGN loc)
+ | Pre_parser.RPAREN loc -> loop (Parser.RPAREN loc)
+ | Pre_parser.SEMICOLON loc -> loop (Parser.SEMICOLON loc)
+ | Pre_parser.SHORT loc -> loop (Parser.SHORT loc)
+ | Pre_parser.SIGNED loc -> loop (Parser.SIGNED loc)
+ | Pre_parser.SIZEOF loc -> loop (Parser.SIZEOF loc)
+ | Pre_parser.SLASH loc -> loop (Parser.SLASH loc)
+ | Pre_parser.STAR loc -> loop (Parser.STAR loc)
+ | Pre_parser.STATIC loc -> loop (Parser.STATIC loc)
+ | Pre_parser.STRING_LITERAL (wide, str, loc) ->
(* Merge consecutive string literals *)
let rec doConcat wide str =
- try
- match Queue.peek tokens with
- | STRING_LITERAL (wide', str', loc) ->
- ignore (Queue.pop tokens);
- let (wide'', str'') = doConcat wide' str' in
- if str'' <> []
- then (wide || wide'', str @ str'')
- else (wide, str)
- | _ ->
- (wide, str)
- with Queue.Empty -> (wide, str) in
- let (wide', str') = doConcat wide str in
- loop STRING_LITERAL't ((wide', str'), loc)
- | STRUCT loc -> loop STRUCT't loc
- | SUB_ASSIGN loc -> loop SUB_ASSIGN't loc
- | SWITCH loc -> loop SWITCH't loc
- | TILDE loc -> loop TILDE't loc
- | TYPEDEF loc -> loop TYPEDEF't loc
- | TYPEDEF_NAME (id, typ, loc)
- | VAR_NAME (id, typ, loc) ->
- let terminal = match !typ with
- | VarId -> VAR_NAME't
- | TypedefId -> TYPEDEF_NAME't
- | OtherId -> OTHER_NAME't
+ match Queue.peek tokens with
+ | Pre_parser.STRING_LITERAL (wide', str', loc) ->
+ ignore (Queue.pop tokens);
+ let (wide'', str'') = doConcat wide' str' in
+ if str'' <> []
+ then (wide || wide'', str @ str'')
+ else (wide, str)
+ | _ -> (wide, str)
+ | exception Queue.Empty -> (wide, str)
in
- loop terminal (id, loc)
- | UNION loc -> loop UNION't loc
- | UNSIGNED loc -> loop UNSIGNED't loc
- | VOID loc -> loop VOID't loc
- | VOLATILE loc -> loop VOLATILE't loc
- | WHILE loc -> loop WHILE't loc
- | XOR_ASSIGN loc -> loop XOR_ASSIGN't loc
- | ALIGNAS loc -> loop ALIGNAS't loc
- | ALIGNOF loc -> loop ALIGNOF't loc
- | ATTRIBUTE loc -> loop ATTRIBUTE't loc
- | ASM loc -> loop ASM't loc
- | PRAGMA (s, loc) -> loop PRAGMA't (s, loc)
- | PRE_NAME _ -> assert false
+ let (wide', str') = doConcat wide str in
+ loop (Parser.STRING_LITERAL ((wide', str'), loc))
+ | Pre_parser.STRUCT loc -> loop (Parser.STRUCT loc)
+ | Pre_parser.SUB_ASSIGN loc -> loop (Parser.SUB_ASSIGN loc)
+ | Pre_parser.SWITCH loc -> loop (Parser.SWITCH loc)
+ | Pre_parser.TILDE loc -> loop (Parser.TILDE loc)
+ | Pre_parser.TYPEDEF loc -> loop (Parser.TYPEDEF loc)
+ | Pre_parser.TYPEDEF_NAME (id, typ, loc)
+ | Pre_parser.VAR_NAME (id, typ, loc) ->
+ begin match !typ with
+ | VarId -> loop (Parser.VAR_NAME (id, loc))
+ | TypedefId -> loop (Parser.TYPEDEF_NAME (id, loc))
+ | OtherId -> loop (Parser.OTHER_NAME (id, loc))
+ end
+ | Pre_parser.UNION loc -> loop (Parser.UNION loc)
+ | Pre_parser.UNSIGNED loc -> loop (Parser.UNSIGNED loc)
+ | Pre_parser.VOID loc -> loop (Parser.VOID loc)
+ | Pre_parser.VOLATILE loc -> loop (Parser.VOLATILE loc)
+ | Pre_parser.WHILE loc -> loop (Parser.WHILE loc)
+ | Pre_parser.XOR_ASSIGN loc -> loop (Parser.XOR_ASSIGN loc)
+ | Pre_parser.ALIGNAS loc -> loop (Parser.ALIGNAS loc)
+ | Pre_parser.ALIGNOF loc -> loop (Parser.ALIGNOF loc)
+ | Pre_parser.ATTRIBUTE loc -> loop (Parser.ATTRIBUTE loc)
+ | Pre_parser.ASM loc -> loop (Parser.ASM loc)
+ | Pre_parser.PRAGMA (s, loc) -> loop (Parser.PRAGMA (s, loc))
+ | Pre_parser.PRE_NAME _ -> assert false
in
- Lazy.from_fun compute_token_stream
+ Lazy.from_fun compute_buffer
}
diff --git a/cparser/MenhirLib/Alphabet.v b/cparser/MenhirLib/Alphabet.v
deleted file mode 100644
index a13f69b0..00000000
--- a/cparser/MenhirLib/Alphabet.v
+++ /dev/null
@@ -1,320 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Jacques-Henri Jourdan, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-Require Import Int31.
-Require Import Cyclic31.
-Require Import Omega.
-Require Import List.
-Require Import Syntax.
-Require Import Relations.
-Require Import RelationClasses.
-
-Local Obligation Tactic := intros.
-
-(** A comparable type is equiped with a [compare] function, that define an order
- relation. **)
-Class Comparable (A:Type) := {
- compare : A -> A -> comparison;
- compare_antisym : forall x y, CompOpp (compare x y) = compare y x;
- compare_trans : forall x y z c,
- (compare x y) = c -> (compare y z) = c -> (compare x z) = c
-}.
-
-Theorem compare_refl {A:Type} (C: Comparable A) :
- forall x, compare x x = Eq.
-Proof.
-intros.
-pose proof (compare_antisym x x).
-destruct (compare x x); intuition; try discriminate.
-Qed.
-
-(** The corresponding order is a strict order. **)
-Definition comparableLt {A:Type} (C: Comparable A) : relation A :=
- fun x y => compare x y = Lt.
-
-Instance ComparableLtStrictOrder {A:Type} (C: Comparable A) :
- StrictOrder (comparableLt C).
-Proof.
-apply Build_StrictOrder.
-unfold Irreflexive, Reflexive, complement, comparableLt.
-intros.
-pose proof H.
-rewrite <- compare_antisym in H.
-rewrite H0 in H.
-discriminate H.
-unfold Transitive, comparableLt.
-intros x y z.
-apply compare_trans.
-Qed.
-
-(** nat is comparable. **)
-Program Instance natComparable : Comparable nat :=
- { compare := Nat.compare }.
-Next Obligation.
-symmetry.
-destruct (Nat.compare x y) as [] eqn:?.
-rewrite Nat.compare_eq_iff in Heqc.
-destruct Heqc.
-rewrite Nat.compare_eq_iff.
-trivial.
-rewrite <- nat_compare_lt in *.
-rewrite <- nat_compare_gt in *.
-trivial.
-rewrite <- nat_compare_lt in *.
-rewrite <- nat_compare_gt in *.
-trivial.
-Qed.
-Next Obligation.
-destruct c.
-rewrite Nat.compare_eq_iff in *; destruct H; assumption.
-rewrite <- nat_compare_lt in *.
-apply (Nat.lt_trans _ _ _ H H0).
-rewrite <- nat_compare_gt in *.
-apply (gt_trans _ _ _ H H0).
-Qed.
-
-(** A pair of comparable is comparable. **)
-Program Instance PairComparable {A:Type} (CA:Comparable A) {B:Type} (CB:Comparable B) :
- Comparable (A*B) :=
- { compare := fun x y =>
- let (xa, xb) := x in let (ya, yb) := y in
- match compare xa ya return comparison with
- | Eq => compare xb yb
- | x => x
- end }.
-Next Obligation.
-destruct x, y.
-rewrite <- (compare_antisym a a0).
-rewrite <- (compare_antisym b b0).
-destruct (compare a a0); intuition.
-Qed.
-Next Obligation.
-destruct x, y, z.
-destruct (compare a a0) as [] eqn:?, (compare a0 a1) as [] eqn:?;
-try (rewrite <- H0 in H; discriminate);
-try (destruct (compare a a1) as [] eqn:?;
- try (rewrite <- compare_antisym in Heqc0;
- rewrite CompOpp_iff in Heqc0;
- rewrite (compare_trans _ _ _ _ Heqc0 Heqc2) in Heqc1;
- discriminate);
- try (rewrite <- compare_antisym in Heqc1;
- rewrite CompOpp_iff in Heqc1;
- rewrite (compare_trans _ _ _ _ Heqc2 Heqc1) in Heqc0;
- discriminate);
- assumption);
-rewrite (compare_trans _ _ _ _ Heqc0 Heqc1);
-try assumption.
-apply (compare_trans _ _ _ _ H H0).
-Qed.
-
-(** Special case of comparable, where equality is usual equality. **)
-Class ComparableUsualEq {A:Type} (C: Comparable A) :=
- compare_eq : forall x y, compare x y = Eq -> x = y.
-
-(** Boolean equality for a [Comparable]. **)
-Definition compare_eqb {A:Type} {C:Comparable A} (x y:A) :=
- match compare x y with
- | Eq => true
- | _ => false
- end.
-
-Theorem compare_eqb_iff {A:Type} {C:Comparable A} {U:ComparableUsualEq C} :
- forall x y, compare_eqb x y = true <-> x = y.
-Proof.
-unfold compare_eqb.
-intuition.
-apply compare_eq.
-destruct (compare x y); intuition; discriminate.
-destruct H.
-rewrite compare_refl; intuition.
-Qed.
-
-(** [Comparable] provides a decidable equality. **)
-Definition compare_eqdec {A:Type} {C:Comparable A} {U:ComparableUsualEq C} (x y:A):
- {x = y} + {x <> y}.
-Proof.
-destruct (compare x y) as [] eqn:?; [left; apply compare_eq; intuition | ..];
- right; intro; destruct H; rewrite compare_refl in Heqc; discriminate.
-Defined.
-
-Instance NComparableUsualEq : ComparableUsualEq natComparable := Nat.compare_eq.
-
-(** A pair of ComparableUsualEq is ComparableUsualEq **)
-Instance PairComparableUsualEq
- {A:Type} {CA:Comparable A} (UA:ComparableUsualEq CA)
- {B:Type} {CB:Comparable B} (UB:ComparableUsualEq CB) :
- ComparableUsualEq (PairComparable CA CB).
-Proof.
-intros x y; destruct x, y; simpl.
-pose proof (compare_eq a a0); pose proof (compare_eq b b0).
-destruct (compare a a0); try discriminate.
-intuition.
-destruct H2, H0.
-reflexivity.
-Qed.
-
-(** An [Finite] type is a type with the list of all elements. **)
-Class Finite (A:Type) := {
- all_list : list A;
- all_list_forall : forall x:A, In x all_list
-}.
-
-(** An alphabet is both [ComparableUsualEq] and [Finite]. **)
-Class Alphabet (A:Type) := {
- AlphabetComparable :> Comparable A;
- AlphabetComparableUsualEq :> ComparableUsualEq AlphabetComparable;
- AlphabetFinite :> Finite A
-}.
-
-(** The [Numbered] class provides a conveniant way to build [Alphabet] instances,
- with a good computationnal complexity. It is mainly a injection from it to
- [Int31] **)
-Class Numbered (A:Type) := {
- inj : A -> int31;
- surj : int31 -> A;
- surj_inj_compat : forall x, surj (inj x) = x;
- inj_bound : int31;
- inj_bound_spec : forall x, (phi (inj x) < phi inj_bound)%Z
-}.
-
-Program Instance NumberedAlphabet {A:Type} (N:Numbered A) : Alphabet A :=
- { AlphabetComparable :=
- {| compare := fun x y => compare31 (inj x) (inj y) |};
- AlphabetFinite :=
- {| all_list := fst (iter_int31 inj_bound _
- (fun p => (cons (surj (snd p)) (fst p), incr (snd p))) ([], 0%int31)) |} }.
-Next Obligation. apply Zcompare_antisym. Qed.
-Next Obligation.
-destruct c. unfold compare31 in *.
-rewrite Z.compare_eq_iff in *. congruence.
-eapply Zcompare_Lt_trans. apply H. apply H0.
-eapply Zcompare_Gt_trans. apply H. apply H0.
-Qed.
-Next Obligation.
-intros x y H. unfold compare, compare31 in H.
-rewrite Z.compare_eq_iff in *.
-rewrite <- surj_inj_compat, <- phi_inv_phi with (inj y), <- H.
-rewrite phi_inv_phi, surj_inj_compat; reflexivity.
-Qed.
-Next Obligation.
-rewrite iter_int31_iter_nat.
-pose proof (inj_bound_spec x).
-match goal with |- In x (fst ?p) => destruct p as [] eqn:? end.
-replace inj_bound with i in H.
-revert l i Heqp x H.
-induction (Z.abs_nat (phi inj_bound)); intros.
-inversion Heqp; clear Heqp; subst.
-rewrite spec_0 in H. pose proof (phi_bounded (inj x)). omega.
-simpl in Heqp.
-destruct nat_rect; specialize (IHn _ _ (eq_refl _) x); simpl in *.
-inversion Heqp. subst. clear Heqp.
-rewrite phi_incr in H.
-pose proof (phi_bounded i0).
-pose proof (phi_bounded (inj x)).
-destruct (Z_lt_le_dec (Z.succ (phi i0)) (2 ^ Z.of_nat size)%Z).
-rewrite Zmod_small in H by omega.
-apply Zlt_succ_le, Zle_lt_or_eq in H.
-destruct H; simpl; auto. left.
-rewrite <- surj_inj_compat, <- phi_inv_phi with (inj x), H, phi_inv_phi; reflexivity.
-replace (Z.succ (phi i0)) with (2 ^ Z.of_nat size)%Z in H by omega.
-rewrite Z_mod_same_full in H.
-exfalso; omega.
-rewrite <- phi_inv_phi with i, <- phi_inv_phi with inj_bound; f_equal.
-pose proof (phi_bounded inj_bound); pose proof (phi_bounded i).
-rewrite <- Z.abs_eq with (phi i), <- Z.abs_eq with (phi inj_bound) by omega.
-clear H H0 H1.
-do 2 rewrite <- Zabs2Nat.id_abs.
-f_equal.
-revert l i Heqp.
-assert (Z.abs_nat (phi inj_bound) < Z.abs_nat (2^31)).
-apply Zabs_nat_lt, phi_bounded.
-induction (Z.abs_nat (phi inj_bound)); intros.
-inversion Heqp; reflexivity.
-inversion Heqp; clear H1 H2 Heqp.
-match goal with |- _ (_ (_ (snd ?p))) = _ => destruct p end.
-pose proof (phi_bounded i0).
-erewrite <- IHn, <- Zabs2Nat.inj_succ in H |- *; eauto; try omega.
-rewrite phi_incr, Zmod_small; intuition; try omega.
-apply inj_lt in H.
-pose proof Z.le_le_succ_r.
-do 2 rewrite Zabs2Nat.id_abs, Z.abs_eq in H; now eauto.
-Qed.
-
-(** Previous class instances for [option A] **)
-Program Instance OptionComparable {A:Type} (C:Comparable A) : Comparable (option A) :=
- { compare := fun x y =>
- match x, y return comparison with
- | None, None => Eq
- | None, Some _ => Lt
- | Some _, None => Gt
- | Some x, Some y => compare x y
- end }.
-Next Obligation.
-destruct x, y; intuition.
-apply compare_antisym.
-Qed.
-Next Obligation.
-destruct x, y, z; try now intuition;
-try (rewrite <- H in H0; discriminate).
-apply (compare_trans _ _ _ _ H H0).
-Qed.
-
-Instance OptionComparableUsualEq {A:Type} {C:Comparable A} (U:ComparableUsualEq C) :
- ComparableUsualEq (OptionComparable C).
-Proof.
-intros x y.
-destruct x, y; intuition; try discriminate.
-rewrite (compare_eq a a0); intuition.
-Qed.
-
-Program Instance OptionFinite {A:Type} (E:Finite A) : Finite (option A) :=
- { all_list := None :: map Some all_list }.
-Next Obligation.
-destruct x; intuition.
-right.
-apply in_map.
-apply all_list_forall.
-Defined.
-
-(** Definitions of [FSet]/[FMap] from [Comparable] **)
-Require Import OrderedTypeAlt.
-Require FSetAVL.
-Require FMapAVL.
-Import OrderedType.
-
-Module Type ComparableM.
- Parameter t : Type.
- Declare Instance tComparable : Comparable t.
-End ComparableM.
-
-Module OrderedTypeAlt_from_ComparableM (C:ComparableM) <: OrderedTypeAlt.
- Definition t := C.t.
- Definition compare : t -> t -> comparison := compare.
-
- Infix "?=" := compare (at level 70, no associativity).
-
- Lemma compare_sym x y : (y?=x) = CompOpp (x?=y).
- Proof. exact (Logic.eq_sym (compare_antisym x y)). Qed.
- Lemma compare_trans c x y z :
- (x?=y) = c -> (y?=z) = c -> (x?=z) = c.
- Proof.
- apply compare_trans.
- Qed.
-End OrderedTypeAlt_from_ComparableM.
-
-Module OrderedType_from_ComparableM (C:ComparableM) <: OrderedType.
- Module Alt := OrderedTypeAlt_from_ComparableM C.
- Include (OrderedType_from_Alt Alt).
-End OrderedType_from_ComparableM.
diff --git a/cparser/MenhirLib/Grammar.v b/cparser/MenhirLib/Grammar.v
deleted file mode 100644
index 8e427cd9..00000000
--- a/cparser/MenhirLib/Grammar.v
+++ /dev/null
@@ -1,166 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Jacques-Henri Jourdan, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-Require Import List.
-Require Import Syntax.
-Require Import Alphabet.
-Require Import Orders.
-Require Tuples.
-
-(** The terminal non-terminal alphabets of the grammar. **)
-Module Type Alphs.
- Parameters terminal nonterminal : Type.
- Declare Instance TerminalAlph: Alphabet terminal.
- Declare Instance NonTerminalAlph: Alphabet nonterminal.
-End Alphs.
-
-(** Definition of the alphabet of symbols, given the alphabet of terminals
- and the alphabet of non terminals **)
-Module Symbol(Import A:Alphs).
-
- Inductive symbol :=
- | T: terminal -> symbol
- | NT: nonterminal -> symbol.
-
- Program Instance SymbolAlph : Alphabet symbol :=
- { AlphabetComparable := {| compare := fun x y =>
- match x, y return comparison with
- | T _, NT _ => Gt
- | NT _, T _ => Lt
- | T x, T y => compare x y
- | NT x, NT y => compare x y
- end |};
- AlphabetFinite := {| all_list :=
- map T all_list++map NT all_list |} }.
- Next Obligation.
- destruct x; destruct y; intuition; apply compare_antisym.
- Qed.
- Next Obligation.
- destruct x; destruct y; destruct z; intuition; try discriminate.
- apply (compare_trans _ t0); intuition.
- apply (compare_trans _ n0); intuition.
- Qed.
- Next Obligation.
- intros x y.
- destruct x; destruct y; try discriminate; intros.
- rewrite (compare_eq t t0); intuition.
- rewrite (compare_eq n n0); intuition.
- Qed.
- Next Obligation.
- rewrite in_app_iff.
- destruct x; [left | right]; apply in_map; apply all_list_forall.
- Qed.
-
-End Symbol.
-
-Module Type T.
- Export Tuples.
-
- Include Alphs <+ Symbol.
-
- (** [symbol_semantic_type] maps a symbols to the type of its semantic
- values. **)
- Parameter symbol_semantic_type: symbol -> Type.
-
- (** The type of productions identifiers **)
- Parameter production : Type.
- Declare Instance ProductionAlph : Alphabet production.
-
- (** Accessors for productions: left hand side, right hand side,
- and semantic action. The semantic actions are given in the form
- of curryfied functions, that take arguments in the reverse order. **)
- Parameter prod_lhs: production -> nonterminal.
- Parameter prod_rhs_rev: production -> list symbol.
- Parameter prod_action:
- forall p:production,
- arrows_left
- (map symbol_semantic_type (rev (prod_rhs_rev p)))
- (symbol_semantic_type (NT (prod_lhs p))).
-
-End T.
-
-Module Defs(Import G:T).
-
- (** A token is a terminal and a semantic value for this terminal. **)
- Definition token := {t:terminal & symbol_semantic_type (T t)}.
-
- (** A grammar creates a relation between word of tokens and semantic values.
- This relation is parametrized by the head symbol. It defines the
- "semantics" of the grammar. This relation is defined by a notion of
- parse tree. **)
- Inductive parse_tree:
- forall (head_symbol:symbol) (word:list token)
- (semantic_value:symbol_semantic_type head_symbol), Type :=
-
- (** A single token has its semantic value as semantic value, for the
- corresponding terminal as head symbol. **)
- | Terminal_pt:
- forall (t:terminal) (sem:symbol_semantic_type (T t)),
- parse_tree (T t)
- [existT (fun t => symbol_semantic_type (T t)) t sem] sem
-
- (** Given a production, if a word has a list of semantic values for the
- right hand side as head symbols, then this word has the semantic value
- given by the semantic action of the production for the left hand side
- as head symbol.**)
- | Non_terminal_pt:
- forall {p:production} {word:list token}
- {semantic_values:tuple (map symbol_semantic_type (rev (prod_rhs_rev p)))},
- parse_tree_list (rev (prod_rhs_rev p)) word semantic_values ->
- parse_tree (NT (prod_lhs p)) word (uncurry (prod_action p) semantic_values)
-
- (** Basically the same relation as before, but for list of head symbols (ie.
- We are building a forest of syntax trees. It is mutually recursive with the
- previous relation **)
- with parse_tree_list:
- forall (head_symbols:list symbol) (word:list token)
- (semantic_values:tuple (map symbol_semantic_type head_symbols)),
- Type :=
-
- (** The empty word has [()] as semantic for [[]] as head symbols list **)
- | Nil_ptl: parse_tree_list [] [] ()
-
- (** The cons of the semantic value for one head symbol and for a list of head
- symbols **)
- | Cons_ptl:
- (** The semantic for the head **)
- forall {head_symbolt:symbol} {wordt:list token}
- {semantic_valuet:symbol_semantic_type head_symbolt},
- parse_tree head_symbolt wordt semantic_valuet ->
-
- (** and the semantic for the tail **)
- forall {head_symbolsq:list symbol} {wordq:list token}
- {semantic_valuesq:tuple (map symbol_semantic_type head_symbolsq)},
- parse_tree_list head_symbolsq wordq semantic_valuesq ->
-
- (** give the semantic of the cons **)
- parse_tree_list
- (head_symbolt::head_symbolsq)
- (wordt++wordq)
- (semantic_valuet, semantic_valuesq).
-
-
- Fixpoint pt_size {head_symbol word sem} (tree:parse_tree head_symbol word sem) :=
- match tree with
- | Terminal_pt _ _ => 1
- | Non_terminal_pt l => S (ptl_size l)
- end
- with ptl_size {head_symbols word sems} (tree:parse_tree_list head_symbols word sems) :=
- match tree with
- | Nil_ptl => 0
- | Cons_ptl t q =>
- pt_size t + ptl_size q
- end.
-End Defs.
diff --git a/cparser/MenhirLib/Interpreter.v b/cparser/MenhirLib/Interpreter.v
deleted file mode 100644
index 4ac02693..00000000
--- a/cparser/MenhirLib/Interpreter.v
+++ /dev/null
@@ -1,228 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Jacques-Henri Jourdan, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-Require Import Streams.
-Require Import List.
-Require Import Syntax.
-Require Automaton.
-Require Import Alphabet.
-
-Module Make(Import A:Automaton.T).
-
-(** The error monad **)
-Inductive result (A:Type) :=
- | Err: result A
- | OK: A -> result A.
-
-Arguments Err [A].
-Arguments OK [A].
-
-Definition bind {A B: Type} (f: result A) (g: A -> result B): result B :=
- match f with
- | OK x => g x
- | Err => Err
- end.
-
-Definition bind2 {A B C: Type} (f: result (A * B)) (g: A -> B -> result C):
- result C :=
- match f with
- | OK (x, y) => g x y
- | Err => Err
- end.
-
-Notation "'do' X <- A ; B" := (bind A (fun X => B))
- (at level 200, X ident, A at level 100, B at level 200).
-
-Notation "'do' ( X , Y ) <- A ; B" := (bind2 A (fun X Y => B))
- (at level 200, X ident, Y ident, A at level 100, B at level 200).
-
-(** Some operations on streams **)
-
-(** Concatenation of a list and a stream **)
-Fixpoint app_str {A:Type} (l:list A) (s:Stream A) :=
- match l with
- | nil => s
- | cons t q => Cons t (app_str q s)
- end.
-
-Infix "++" := app_str (right associativity, at level 60).
-
-Lemma app_str_app_assoc {A:Type} (l1 l2:list A) (s:Stream A) :
- l1 ++ (l2 ++ s) = (l1 ++ l2) ++ s.
-Proof.
-induction l1.
-reflexivity.
-simpl.
-rewrite IHl1.
-reflexivity.
-Qed.
-
-(** The type of a non initial state: the type of semantic values associated
- with the last symbol of this state. *)
-Definition noninitstate_type state :=
- symbol_semantic_type (last_symb_of_non_init_state state).
-
-(** The stack of the automaton : it can be either nil or contains a non
- initial state, a semantic value for the symbol associted with this state,
- and a nested stack. **)
-Definition stack := list (sigT noninitstate_type). (* eg. list {state & state_type state} *)
-
-Section Init.
-
-Variable init : initstate.
-
-(** The top state of a stack **)
-Definition state_of_stack (stack:stack): state :=
- match stack with
- | [] => init
- | existT _ s _::_ => s
- end.
-
-(** [pop] pops some symbols from the stack. It returns the popped semantic
- values using [sem_popped] as an accumulator and discards the popped
- states.**)
-Fixpoint pop (symbols_to_pop:list symbol) (stack_cur:stack):
- forall {A:Type} (action:arrows_right A (map symbol_semantic_type symbols_to_pop)),
- result (stack * A) :=
- match symbols_to_pop return forall {A:Type} (action:arrows_right A (map _ symbols_to_pop)), result (stack * A) with
- | [] => fun A action => OK (stack_cur, action)
- | t::q => fun A action =>
- match stack_cur with
- | existT _ state_cur sem::stack_rec =>
- match compare_eqdec (last_symb_of_non_init_state state_cur) t with
- | left e =>
- let sem_conv := eq_rect _ symbol_semantic_type sem _ e in
- pop q stack_rec (action sem_conv)
- | right _ => Err
- end
- | [] => Err
- end
- end.
-
-(** [step_result] represents the result of one step of the automaton : it can
- fail, accept or progress. [Fail_sr] means that the input is incorrect.
- [Accept_sr] means that this is the last step of the automaton, and it
- returns the semantic value of the input word. [Progress_sr] means that
- some progress has been made, but new steps are needed in order to accept
- a word.
-
- For [Accept_sr] and [Progress_sr], the result contains the new input buffer.
-
- [Fail_sr] means that the input word is rejected by the automaton. It is
- different to [Err] (from the error monad), which mean that the automaton is
- bogus and has perfomed a forbidden action. **)
-Inductive step_result :=
- | Fail_sr: step_result
- | Accept_sr: symbol_semantic_type (NT (start_nt init)) -> Stream token -> step_result
- | Progress_sr: stack -> Stream token -> step_result.
-
-Program Definition prod_action':
- forall p,
- arrows_right (symbol_semantic_type (NT (prod_lhs p)))
- (map symbol_semantic_type (prod_rhs_rev p)):=
- fun p =>
- eq_rect _ (fun x => x) (prod_action p) _ _.
-Next Obligation.
-unfold arrows_left, arrows_right; simpl.
-rewrite <- fold_left_rev_right, <- map_rev, rev_involutive.
-reflexivity.
-Qed.
-
-(** [reduce_step] does a reduce action :
- - pops some elements from the stack
- - execute the action of the production
- - follows the goto for the produced non terminal symbol **)
-Definition reduce_step stack_cur production buffer: result step_result :=
- do (stack_new, sem) <-
- pop (prod_rhs_rev production) stack_cur (prod_action' production);
- match goto_table (state_of_stack stack_new) (prod_lhs production) with
- | Some (exist _ state_new e) =>
- let sem := eq_rect _ _ sem _ e in
- OK (Progress_sr (existT noninitstate_type state_new sem::stack_new) buffer)
- | None =>
- match stack_new with
- | [] =>
- match compare_eqdec (prod_lhs production) (start_nt init) with
- | left e =>
- let sem := eq_rect _ (fun nt => symbol_semantic_type (NT nt)) sem _ e in
- OK (Accept_sr sem buffer)
- | right _ => Err
- end
- | _::_ => Err
- end
- end.
-
-(** One step of parsing. **)
-Definition step stack_cur buffer: result step_result :=
- match action_table (state_of_stack stack_cur) with
- | Default_reduce_act production =>
- reduce_step stack_cur production buffer
- | Lookahead_act awt =>
- match Streams.hd buffer with
- | existT _ term sem =>
- match awt term with
- | Shift_act state_new e =>
- let sem_conv := eq_rect _ symbol_semantic_type sem _ e in
- OK (Progress_sr (existT noninitstate_type state_new sem_conv::stack_cur)
- (Streams.tl buffer))
- | Reduce_act production =>
- reduce_step stack_cur production buffer
- | Fail_action =>
- OK Fail_sr
- end
- end
- end.
-
-(** The parsing use a [nat] parameter [n_steps], so that we do not have to prove
- terminaison, which is difficult. So the result of a parsing is either
- a failure (the automaton has rejected the input word), either a timeout
- (the automaton has spent all the given [n_steps]), either a parsed semantic
- value with a rest of the input buffer.
-**)
-Inductive parse_result :=
- | Fail_pr: parse_result
- | Timeout_pr: parse_result
- | Parsed_pr: symbol_semantic_type (NT (start_nt init)) -> Stream token -> parse_result.
-
-Fixpoint parse_fix stack_cur buffer n_steps: result parse_result:=
- match n_steps with
- | O => OK Timeout_pr
- | S it =>
- do r <- step stack_cur buffer;
- match r with
- | Fail_sr => OK Fail_pr
- | Accept_sr t buffer_new => OK (Parsed_pr t buffer_new)
- | Progress_sr s buffer_new => parse_fix s buffer_new it
- end
- end.
-
-Definition parse buffer n_steps: result parse_result :=
- parse_fix [] buffer n_steps.
-
-End Init.
-
-Arguments Fail_sr [init].
-Arguments Accept_sr [init] _ _.
-Arguments Progress_sr [init] _ _.
-
-Arguments Fail_pr [init].
-Arguments Timeout_pr [init].
-Arguments Parsed_pr [init] _ _.
-
-End Make.
-
-Module Type T(A:Automaton.T).
- Include (Make A).
-End T.
diff --git a/cparser/MenhirLib/Interpreter_complete.v b/cparser/MenhirLib/Interpreter_complete.v
deleted file mode 100644
index 2e64b8da..00000000
--- a/cparser/MenhirLib/Interpreter_complete.v
+++ /dev/null
@@ -1,686 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Jacques-Henri Jourdan, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-Require Import Streams.
-Require Import ProofIrrelevance.
-Require Import Equality.
-Require Import List.
-Require Import Syntax.
-Require Import Alphabet.
-Require Import Arith.
-Require Grammar.
-Require Automaton.
-Require Interpreter.
-Require Validator_complete.
-
-Module Make(Import A:Automaton.T) (Import Inter:Interpreter.T A).
-Module Import Valid := Validator_complete.Make A.
-
-(** * Completeness Proof **)
-
-Section Completeness_Proof.
-
-Hypothesis complete: complete.
-
-Proposition nullable_stable: nullable_stable.
-Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
-Proposition first_stable: first_stable.
-Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
-Proposition start_future: start_future.
-Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
-Proposition terminal_shift: terminal_shift.
-Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
-Proposition end_reduce: end_reduce.
-Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
-Proposition start_goto: start_goto.
-Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
-Proposition non_terminal_goto: non_terminal_goto.
-Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
-Proposition non_terminal_closed: non_terminal_closed.
-Proof. pose proof complete; unfold Valid.complete in H; intuition. Qed.
-
-(** If the nullable predicate has been validated, then it is correct. **)
-Lemma nullable_correct:
- forall head sem word, word = [] ->
- parse_tree head word sem -> nullable_symb head = true
-with nullable_correct_list:
- forall heads sems word, word = [] ->
- parse_tree_list heads word sems -> nullable_word heads = true.
-Proof with eauto.
-intros.
-destruct X.
-congruence.
-apply nullable_stable...
-intros.
-destruct X; simpl...
-apply andb_true_intro.
-apply app_eq_nil in H; destruct H; split...
-Qed.
-
-(** If the first predicate has been validated, then it is correct. **)
-Lemma first_correct:
- forall head sem word t q, word = t::q ->
- parse_tree head word sem ->
- TerminalSet.In (projT1 t) (first_symb_set head)
-with first_correct_list:
- forall heads sems word t q, word = t::q ->
- parse_tree_list heads word sems ->
- TerminalSet.In (projT1 t) (first_word_set heads).
-Proof with eauto.
-intros.
-destruct X.
-inversion H; subst.
-apply TerminalSet.singleton_2, compare_refl...
-apply first_stable...
-intros.
-destruct X.
-congruence.
-simpl.
-case_eq wordt; intros.
-erewrite nullable_correct...
-apply TerminalSet.union_3.
-subst...
-rewrite H0 in *; inversion H; destruct H2.
-destruct (nullable_symb head_symbolt)...
-apply TerminalSet.union_2...
-Qed.
-
-Variable init: initstate.
-Variable full_word: list token.
-Variable buffer_end: Stream token.
-Variable full_sem: symbol_semantic_type (NT (start_nt init)).
-
-Inductive pt_zipper:
- forall (hole_symb:symbol) (hole_word:list token)
- (hole_sem:symbol_semantic_type hole_symb), Type :=
-| Top_ptz:
- pt_zipper (NT (start_nt init)) (full_word) (full_sem)
-| Cons_ptl_ptz:
- forall {head_symbolt:symbol}
- {wordt:list token}
- {semantic_valuet:symbol_semantic_type head_symbolt},
-
- forall {head_symbolsq:list symbol}
- {wordq:list token}
- {semantic_valuesq:tuple (map symbol_semantic_type head_symbolsq)},
- parse_tree_list head_symbolsq wordq semantic_valuesq ->
-
- ptl_zipper (head_symbolt::head_symbolsq) (wordt++wordq)
- (semantic_valuet,semantic_valuesq) ->
-
- pt_zipper head_symbolt wordt semantic_valuet
-with ptl_zipper:
- forall (hole_symbs:list symbol) (hole_word:list token)
- (hole_sems:tuple (map symbol_semantic_type hole_symbs)), Type :=
-| Non_terminal_pt_ptlz:
- forall {p:production} {word:list token}
- {semantic_values:tuple (map symbol_semantic_type (rev (prod_rhs_rev p)))},
- pt_zipper (NT (prod_lhs p)) word (uncurry (prod_action p) semantic_values) ->
- ptl_zipper (rev (prod_rhs_rev p)) word semantic_values
-
-| Cons_ptl_ptlz:
- forall {head_symbolt:symbol}
- {wordt:list token}
- {semantic_valuet:symbol_semantic_type head_symbolt},
- parse_tree head_symbolt wordt semantic_valuet ->
-
- forall {head_symbolsq:list symbol}
- {wordq:list token}
- {semantic_valuesq:tuple (map symbol_semantic_type head_symbolsq)},
-
- ptl_zipper (head_symbolt::head_symbolsq) (wordt++wordq)
- (semantic_valuet,semantic_valuesq) ->
-
- ptl_zipper head_symbolsq wordq semantic_valuesq.
-
-Fixpoint ptlz_cost {hole_symbs hole_word hole_sems}
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems) :=
- match ptlz with
- | Non_terminal_pt_ptlz ptz =>
- ptz_cost ptz
- | Cons_ptl_ptlz pt ptlz' =>
- ptlz_cost ptlz'
- end
-with ptz_cost {hole_symb hole_word hole_sem}
- (ptz:pt_zipper hole_symb hole_word hole_sem) :=
- match ptz with
- | Top_ptz => 0
- | Cons_ptl_ptz ptl ptlz' =>
- 1 + ptl_size ptl + ptlz_cost ptlz'
- end.
-
-Inductive pt_dot: Type :=
-| Reduce_ptd: ptl_zipper [] [] () -> pt_dot
-| Shift_ptd:
- forall (term:terminal) (sem: symbol_semantic_type (T term))
- {symbolsq wordq semsq},
- parse_tree_list symbolsq wordq semsq ->
- ptl_zipper (T term::symbolsq) (existT (fun t => symbol_semantic_type (T t)) term sem::wordq) (sem, semsq) ->
- pt_dot.
-
-Definition ptd_cost (ptd:pt_dot) :=
- match ptd with
- | Reduce_ptd ptlz => ptlz_cost ptlz
- | Shift_ptd _ _ ptl ptlz => 1 + ptl_size ptl + ptlz_cost ptlz
- end.
-
-Fixpoint ptlz_buffer {hole_symbs hole_word hole_sems}
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems): Stream token :=
- match ptlz with
- | Non_terminal_pt_ptlz ptz =>
- ptz_buffer ptz
- | Cons_ptl_ptlz _ ptlz' =>
- ptlz_buffer ptlz'
- end
-with ptz_buffer {hole_symb hole_word hole_sem}
- (ptz:pt_zipper hole_symb hole_word hole_sem): Stream token :=
- match ptz with
- | Top_ptz => buffer_end
- | @Cons_ptl_ptz _ _ _ _ wordq _ ptl ptlz' =>
- wordq++ptlz_buffer ptlz'
- end.
-
-Definition ptd_buffer (ptd:pt_dot) :=
- match ptd with
- | Reduce_ptd ptlz => ptlz_buffer ptlz
- | @Shift_ptd term sem _ wordq _ _ ptlz =>
- Cons (existT (fun t => symbol_semantic_type (T t)) term sem)
- (wordq ++ ptlz_buffer ptlz)
- end.
-
-Fixpoint ptlz_prod {hole_symbs hole_word hole_sems}
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems): production :=
- match ptlz with
- | @Non_terminal_pt_ptlz prod _ _ _ => prod
- | Cons_ptl_ptlz _ ptlz' =>
- ptlz_prod ptlz'
- end.
-
-Fixpoint ptlz_past {hole_symbs hole_word hole_sems}
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems): list symbol :=
- match ptlz with
- | Non_terminal_pt_ptlz _ => []
- | @Cons_ptl_ptlz s _ _ _ _ _ _ ptlz' => s::ptlz_past ptlz'
- end.
-
-Lemma ptlz_past_ptlz_prod:
- forall hole_symbs hole_word hole_sems
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems),
- rev_append hole_symbs (ptlz_past ptlz) = prod_rhs_rev (ptlz_prod ptlz).
-Proof.
-fix ptlz_past_ptlz_prod 4.
-destruct ptlz; simpl.
-rewrite <- rev_alt, rev_involutive; reflexivity.
-apply (ptlz_past_ptlz_prod _ _ _ ptlz).
-Qed.
-
-Definition ptlz_state_compat {hole_symbs hole_word hole_sems}
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems)
- (state:state): Prop :=
- state_has_future state (ptlz_prod ptlz) hole_symbs
- (projT1 (Streams.hd (ptlz_buffer ptlz))).
-
-Fixpoint ptlz_stack_compat {hole_symbs hole_word hole_sems}
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems)
- (stack:stack): Prop :=
- ptlz_state_compat ptlz (state_of_stack init stack) /\
- match ptlz with
- | Non_terminal_pt_ptlz ptz =>
- ptz_stack_compat ptz stack
- | @Cons_ptl_ptlz _ _ sem _ _ _ _ ptlz' =>
- match stack with
- | [] => False
- | existT _ _ sem'::stackq =>
- ptlz_stack_compat ptlz' stackq /\
- sem ~= sem'
- end
- end
-with ptz_stack_compat {hole_symb hole_word hole_sem}
- (ptz:pt_zipper hole_symb hole_word hole_sem)
- (stack:stack): Prop :=
- match ptz with
- | Top_ptz => stack = []
- | Cons_ptl_ptz _ ptlz' =>
- ptlz_stack_compat ptlz' stack
- end.
-
-Lemma ptlz_stack_compat_ptlz_state_compat:
- forall hole_symbs hole_word hole_sems
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems)
- (stack:stack),
- ptlz_stack_compat ptlz stack -> ptlz_state_compat ptlz (state_of_stack init stack).
-Proof.
-destruct ptlz; simpl; intuition.
-Qed.
-
-Definition ptd_stack_compat (ptd:pt_dot) (stack:stack): Prop :=
- match ptd with
- | Reduce_ptd ptlz => ptlz_stack_compat ptlz stack
- | Shift_ptd _ _ _ ptlz => ptlz_stack_compat ptlz stack
- end.
-
-Fixpoint build_pt_dot {hole_symbs hole_word hole_sems}
- (ptl:parse_tree_list hole_symbs hole_word hole_sems)
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems)
- :pt_dot :=
- match ptl in parse_tree_list hole_symbs hole_word hole_sems
- return ptl_zipper hole_symbs hole_word hole_sems -> _
- with
- | Nil_ptl => fun ptlz =>
- Reduce_ptd ptlz
- | Cons_ptl pt ptl' =>
- match pt in parse_tree hole_symb hole_word hole_sem
- return ptl_zipper (hole_symb::_) (hole_word++_) (hole_sem,_) -> _
- with
- | Terminal_pt term sem => fun ptlz =>
- Shift_ptd term sem ptl' ptlz
- | Non_terminal_pt ptl'' => fun ptlz =>
- build_pt_dot ptl''
- (Non_terminal_pt_ptlz (Cons_ptl_ptz ptl' ptlz))
- end
- end ptlz.
-
-Lemma build_pt_dot_cost:
- forall hole_symbs hole_word hole_sems
- (ptl:parse_tree_list hole_symbs hole_word hole_sems)
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems),
- ptd_cost (build_pt_dot ptl ptlz) = ptl_size ptl + ptlz_cost ptlz.
-Proof.
-fix build_pt_dot_cost 4.
-destruct ptl; intros.
-reflexivity.
-destruct p.
-reflexivity.
-simpl; rewrite build_pt_dot_cost.
-simpl; rewrite <- plus_n_Sm, Nat.add_assoc; reflexivity.
-Qed.
-
-Lemma build_pt_dot_buffer:
- forall hole_symbs hole_word hole_sems
- (ptl:parse_tree_list hole_symbs hole_word hole_sems)
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems),
- ptd_buffer (build_pt_dot ptl ptlz) = hole_word ++ ptlz_buffer ptlz.
-Proof.
-fix build_pt_dot_buffer 4.
-destruct ptl; intros.
-reflexivity.
-destruct p.
-reflexivity.
-simpl; rewrite build_pt_dot_buffer.
-apply app_str_app_assoc.
-Qed.
-
-Lemma ptd_stack_compat_build_pt_dot:
- forall hole_symbs hole_word hole_sems
- (ptl:parse_tree_list hole_symbs hole_word hole_sems)
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems)
- (stack:stack),
- ptlz_stack_compat ptlz stack ->
- ptd_stack_compat (build_pt_dot ptl ptlz) stack.
-Proof.
-fix ptd_stack_compat_build_pt_dot 4.
-destruct ptl; intros.
-eauto.
-destruct p.
-eauto.
-simpl.
-apply ptd_stack_compat_build_pt_dot.
-split.
-apply ptlz_stack_compat_ptlz_state_compat, non_terminal_closed in H.
-apply H; clear H; eauto.
-destruct wordq.
-right; split.
-eauto.
-eapply nullable_correct_list; eauto.
-left.
-eapply first_correct_list; eauto.
-eauto.
-Qed.
-
-Program Fixpoint pop_ptlz {hole_symbs hole_word hole_sems}
- (ptl:parse_tree_list hole_symbs hole_word hole_sems)
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems):
- { word:_ & { sem:_ &
- (pt_zipper (NT (prod_lhs (ptlz_prod ptlz))) word sem *
- parse_tree (NT (prod_lhs (ptlz_prod ptlz))) word sem)%type } } :=
- match ptlz in ptl_zipper hole_symbs hole_word hole_sems
- return parse_tree_list hole_symbs hole_word hole_sems ->
- { word:_ & { sem:_ &
- (pt_zipper (NT (prod_lhs (ptlz_prod ptlz))) word sem *
- parse_tree (NT (prod_lhs (ptlz_prod ptlz))) word sem)%type } }
- with
- | @Non_terminal_pt_ptlz prod word sem ptz => fun ptl =>
- let sem := uncurry (prod_action prod) sem in
- existT _ word (existT _ sem (ptz, Non_terminal_pt ptl))
- | Cons_ptl_ptlz pt ptlz' => fun ptl =>
- pop_ptlz (Cons_ptl pt ptl) ptlz'
- end ptl.
-
-Lemma pop_ptlz_cost:
- forall hole_symbs hole_word hole_sems
- (ptl:parse_tree_list hole_symbs hole_word hole_sems)
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems),
- let 'existT _ word (existT _ sem (ptz, pt)) := pop_ptlz ptl ptlz in
- ptlz_cost ptlz = ptz_cost ptz.
-Proof.
-fix pop_ptlz_cost 5.
-destruct ptlz.
-reflexivity.
-simpl; apply pop_ptlz_cost.
-Qed.
-
-Lemma pop_ptlz_buffer:
- forall hole_symbs hole_word hole_sems
- (ptl:parse_tree_list hole_symbs hole_word hole_sems)
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems),
- let 'existT _ word (existT _ sem (ptz, pt)) := pop_ptlz ptl ptlz in
- ptlz_buffer ptlz = ptz_buffer ptz.
-Proof.
-fix pop_ptlz_buffer 5.
-destruct ptlz.
-reflexivity.
-simpl; apply pop_ptlz_buffer.
-Qed.
-
-Lemma pop_ptlz_pop_stack_compat_converter:
- forall A hole_symbs hole_word hole_sems (ptlz:ptl_zipper hole_symbs hole_word hole_sems),
- arrows_left (map symbol_semantic_type (rev (prod_rhs_rev (ptlz_prod ptlz)))) A =
- arrows_left (map symbol_semantic_type hole_symbs)
- (arrows_right A (map symbol_semantic_type (ptlz_past ptlz))).
-Proof.
-intros.
-rewrite <- ptlz_past_ptlz_prod.
-unfold arrows_right, arrows_left.
-rewrite rev_append_rev, map_rev, map_app, map_rev, <- fold_left_rev_right, rev_involutive, fold_right_app.
-rewrite fold_left_rev_right.
-reflexivity.
-Qed.
-
-Lemma pop_ptlz_pop_stack_compat:
- forall hole_symbs hole_word hole_sems
- (ptl:parse_tree_list hole_symbs hole_word hole_sems)
- (ptlz:ptl_zipper hole_symbs hole_word hole_sems)
- (stack:stack),
-
- ptlz_stack_compat ptlz stack ->
-
- let action' :=
- eq_rect _ (fun x=>x) (prod_action (ptlz_prod ptlz)) _
- (pop_ptlz_pop_stack_compat_converter _ _ _ _ _)
- in
- let 'existT _ word (existT _ sem (ptz, pt)) := pop_ptlz ptl ptlz in
- match pop (ptlz_past ptlz) stack (uncurry action' hole_sems) with
- | OK (stack', sem') =>
- ptz_stack_compat ptz stack' /\ sem = sem'
- | Err => True
- end.
-Proof.
-Opaque AlphabetComparable AlphabetComparableUsualEq.
-fix pop_ptlz_pop_stack_compat 5.
-destruct ptlz. intros; simpl.
-split.
-apply H.
-eapply (f_equal (fun X => uncurry X semantic_values)).
-apply JMeq_eq, JMeq_sym, JMeq_eqrect with (P:=fun x => x).
-simpl; intros; destruct stack0.
-destruct (proj2 H).
-simpl in H; destruct H.
-destruct s as (state, sem').
-destruct H0.
-specialize (pop_ptlz_pop_stack_compat _ _ _ (Cons_ptl p ptl) ptlz _ H0).
-destruct (pop_ptlz (Cons_ptl p ptl) ptlz) as [word [sem []]].
-destruct (compare_eqdec (last_symb_of_non_init_state state) head_symbolt); intuition.
-eapply JMeq_sym, JMeq_trans, JMeq_sym, JMeq_eq in H1; [|apply JMeq_eqrect with (e:=e)].
-rewrite <- H1.
-simpl in pop_ptlz_pop_stack_compat.
-erewrite proof_irrelevance with (p1:=pop_ptlz_pop_stack_compat_converter _ _ _ _ _).
-apply pop_ptlz_pop_stack_compat.
-Transparent AlphabetComparable AlphabetComparableUsualEq.
-Qed.
-
-Definition next_ptd (ptd:pt_dot): option pt_dot :=
- match ptd with
- | Shift_ptd term sem ptl ptlz =>
- Some (build_pt_dot ptl (Cons_ptl_ptlz (Terminal_pt term sem) ptlz))
- | Reduce_ptd ptlz =>
- let 'existT _ _ (existT _ _ (ptz, pt)) := pop_ptlz Nil_ptl ptlz in
- match ptz in pt_zipper sym _ _ return parse_tree sym _ _ -> _ with
- | Top_ptz => fun pt => None
- | Cons_ptl_ptz ptl ptlz' =>
- fun pt => Some (build_pt_dot ptl (Cons_ptl_ptlz pt ptlz'))
- end pt
- end.
-
-Lemma next_ptd_cost:
- forall ptd,
- match next_ptd ptd with
- | None => ptd_cost ptd = 0
- | Some ptd' => ptd_cost ptd = S (ptd_cost ptd')
- end.
-Proof.
-destruct ptd. unfold next_ptd.
-pose proof (pop_ptlz_cost _ _ _ Nil_ptl p).
-destruct (pop_ptlz Nil_ptl p) as [word[sem[[]]]].
-assumption.
-rewrite build_pt_dot_cost.
-assumption.
-simpl; rewrite build_pt_dot_cost; reflexivity.
-Qed.
-
-Lemma reduce_step_next_ptd:
- forall (ptlz:ptl_zipper [] [] ()) (stack:stack),
- ptlz_stack_compat ptlz stack ->
- match reduce_step init stack (ptlz_prod ptlz) (ptlz_buffer ptlz) with
- | OK Fail_sr =>
- False
- | OK (Accept_sr sem buffer) =>
- sem = full_sem /\ buffer = buffer_end /\ next_ptd (Reduce_ptd ptlz) = None
- | OK (Progress_sr stack buffer) =>
- match next_ptd (Reduce_ptd ptlz) with
- | None => False
- | Some ptd =>
- ptd_stack_compat ptd stack /\ buffer = ptd_buffer ptd
- end
- | Err =>
- True
- end.
-Proof.
-intros.
-unfold reduce_step, next_ptd.
-apply pop_ptlz_pop_stack_compat with (ptl:=Nil_ptl) in H.
-pose proof (pop_ptlz_buffer _ _ _ Nil_ptl ptlz).
-destruct (pop_ptlz Nil_ptl ptlz) as [word [sem [ptz pt]]].
-rewrite H0; clear H0.
-revert H.
-match goal with
- |- match ?p1 with Err => _ | OK _ => _ end -> match bind2 ?p2 _ with Err => _ | OK _ => _ end =>
- replace p1 with p2; [destruct p2 as [|[]]; intros|]
-end.
-assumption.
-simpl.
-destruct H; subst.
-generalize dependent s0.
-generalize (prod_lhs (ptlz_prod ptlz)); clear ptlz stack0.
-dependent destruction ptz; intros.
-simpl in H; subst; simpl.
-pose proof start_goto; unfold Valid.start_goto in H; rewrite H.
-destruct (compare_eqdec (start_nt init) (start_nt init)); intuition.
-apply JMeq_eq, JMeq_eqrect with (P:=fun nt => symbol_semantic_type (NT nt)).
-pose proof (ptlz_stack_compat_ptlz_state_compat _ _ _ _ _ H).
-apply non_terminal_goto in H0.
-destruct (goto_table (state_of_stack init s) n) as [[]|]; intuition.
-apply ptd_stack_compat_build_pt_dot; simpl; intuition.
-symmetry; apply JMeq_eqrect with (P:=symbol_semantic_type).
-symmetry; apply build_pt_dot_buffer.
-destruct s; intuition.
-simpl in H; apply ptlz_stack_compat_ptlz_state_compat in H.
-destruct (H0 _ _ _ H eq_refl).
-generalize (pop_ptlz_pop_stack_compat_converter (symbol_semantic_type (NT (prod_lhs (ptlz_prod ptlz)))) _ _ _ ptlz).
-pose proof (ptlz_past_ptlz_prod _ _ _ ptlz); simpl in H.
-rewrite H; clear H.
-intro; f_equal; simpl.
-apply JMeq_eq.
-etransitivity.
-apply JMeq_eqrect with (P:=fun x => x).
-symmetry.
-apply JMeq_eqrect with (P:=fun x => x).
-Qed.
-
-Lemma step_next_ptd:
- forall (ptd:pt_dot) (stack:stack),
- ptd_stack_compat ptd stack ->
- match step init stack (ptd_buffer ptd) with
- | OK Fail_sr =>
- False
- | OK (Accept_sr sem buffer) =>
- sem = full_sem /\ buffer = buffer_end /\ next_ptd ptd = None
- | OK (Progress_sr stack buffer) =>
- match next_ptd ptd with
- | None => False
- | Some ptd =>
- ptd_stack_compat ptd stack /\ buffer = ptd_buffer ptd
- end
- | Err =>
- True
- end.
-Proof.
-intros.
-destruct ptd.
-pose proof (ptlz_stack_compat_ptlz_state_compat _ _ _ _ _ H).
-apply end_reduce in H0.
-unfold step.
-destruct (action_table (state_of_stack init stack0)).
-rewrite H0 by reflexivity.
-apply reduce_step_next_ptd; assumption.
-simpl; destruct (Streams.hd (ptlz_buffer p)); simpl in H0.
-destruct (l x); intuition; rewrite H1.
-apply reduce_step_next_ptd; assumption.
-pose proof (ptlz_stack_compat_ptlz_state_compat _ _ _ _ _ H).
-apply terminal_shift in H0.
-unfold step.
-destruct (action_table (state_of_stack init stack0)); intuition.
-simpl; destruct (Streams.hd (ptlz_buffer p0)) as [] eqn:?.
-destruct (l term); intuition.
-apply ptd_stack_compat_build_pt_dot; simpl; intuition.
-unfold ptlz_state_compat; simpl; destruct Heqt; assumption.
-symmetry; apply JMeq_eqrect with (P:=symbol_semantic_type).
-rewrite build_pt_dot_buffer; reflexivity.
-Qed.
-
-Lemma parse_fix_complete:
- forall (ptd:pt_dot) (stack:stack) (n_steps:nat),
- ptd_stack_compat ptd stack ->
- match parse_fix init stack (ptd_buffer ptd) n_steps with
- | OK (Parsed_pr sem_res buffer_end_res) =>
- sem_res = full_sem /\ buffer_end_res = buffer_end /\
- S (ptd_cost ptd) <= n_steps
- | OK Fail_pr => False
- | OK Timeout_pr => n_steps < S (ptd_cost ptd)
- | Err => True
- end.
-Proof.
-fix parse_fix_complete 3.
-destruct n_steps; intros; simpl.
-apply Nat.lt_0_succ.
-apply step_next_ptd in H.
-pose proof (next_ptd_cost ptd).
-destruct (step init stack0 (ptd_buffer ptd)) as [|[]]; simpl; intuition.
-rewrite H3 in H0; rewrite H0.
-apply le_n_S, Nat.le_0_l.
-destruct (next_ptd ptd); intuition; subst.
-eapply parse_fix_complete with (n_steps:=n_steps) in H1.
-rewrite H0.
-destruct (parse_fix init s (ptd_buffer p) n_steps) as [|[]]; try assumption.
-apply lt_n_S; assumption.
-destruct H1 as [H1 []]; split; [|split]; try assumption.
-apply le_n_S; assumption.
-Qed.
-
-Variable full_pt: parse_tree (NT (start_nt init)) full_word full_sem.
-
-Definition init_ptd :=
- match full_pt in parse_tree head full_word full_sem return
- pt_zipper head full_word full_sem ->
- match head return Type with | T _ => unit | NT _ => pt_dot end
- with
- | Terminal_pt _ _ => fun _ => ()
- | Non_terminal_pt ptl =>
- fun top => build_pt_dot ptl (Non_terminal_pt_ptlz top)
- end Top_ptz.
-
-Lemma init_ptd_compat:
- ptd_stack_compat init_ptd [].
-Proof.
-unfold init_ptd.
-assert (ptz_stack_compat Top_ptz []) by reflexivity.
-pose proof (start_future init); revert H0.
-generalize dependent Top_ptz.
-generalize dependent full_word.
-generalize full_sem.
-generalize (start_nt init).
-dependent destruction full_pt0.
-intros.
-apply ptd_stack_compat_build_pt_dot; simpl; intuition.
-apply H0; reflexivity.
-Qed.
-
-Lemma init_ptd_cost:
- S (ptd_cost init_ptd) = pt_size full_pt.
-Proof.
-unfold init_ptd.
-assert (ptz_cost Top_ptz = 0) by reflexivity.
-generalize dependent Top_ptz.
-generalize dependent full_word.
-generalize full_sem.
-generalize (start_nt init).
-dependent destruction full_pt0.
-intros.
-rewrite build_pt_dot_cost; simpl.
-rewrite H, Nat.add_0_r; reflexivity.
-Qed.
-
-Lemma init_ptd_buffer:
- ptd_buffer init_ptd = full_word ++ buffer_end.
-Proof.
-unfold init_ptd.
-assert (ptz_buffer Top_ptz = buffer_end) by reflexivity.
-generalize dependent Top_ptz.
-generalize dependent full_word.
-generalize full_sem.
-generalize (start_nt init).
-dependent destruction full_pt0.
-intros.
-rewrite build_pt_dot_buffer; simpl.
-rewrite H; reflexivity.
-Qed.
-
-Theorem parse_complete n_steps:
- match parse init (full_word ++ buffer_end) n_steps with
- | OK (Parsed_pr sem_res buffer_end_res) =>
- sem_res = full_sem /\ buffer_end_res = buffer_end /\
- pt_size full_pt <= n_steps
- | OK Fail_pr => False
- | OK Timeout_pr => n_steps < pt_size full_pt
- | Err => True
- end.
-Proof.
-pose proof (parse_fix_complete init_ptd [] n_steps init_ptd_compat).
-rewrite init_ptd_buffer, init_ptd_cost in H.
-apply H.
-Qed.
-
-End Completeness_Proof.
-
-End Make.
diff --git a/cparser/MenhirLib/Interpreter_correct.v b/cparser/MenhirLib/Interpreter_correct.v
deleted file mode 100644
index 1263d4e3..00000000
--- a/cparser/MenhirLib/Interpreter_correct.v
+++ /dev/null
@@ -1,228 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Jacques-Henri Jourdan, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-Require Import Streams.
-Require Import List.
-Require Import Syntax.
-Require Import Equality.
-Require Import Alphabet.
-Require Grammar.
-Require Automaton.
-Require Interpreter.
-
-Module Make(Import A:Automaton.T) (Import Inter:Interpreter.T A).
-
-(** * Correctness of the interpreter **)
-
-(** We prove that, in any case, if the interpreter accepts returning a
- semantic value, then this is a semantic value of the input **)
-
-Section Init.
-
-Variable init:initstate.
-
-(** [word_has_stack_semantics] relates a word with a state, stating that the
- word is a concatenation of words that have the semantic values stored in
- the stack. **)
-Inductive word_has_stack_semantics:
- forall (word:list token) (stack:stack), Prop :=
- | Nil_stack_whss: word_has_stack_semantics [] []
- | Cons_stack_whss:
- forall (wordq:list token) (stackq:stack),
- word_has_stack_semantics wordq stackq ->
-
- forall (wordt:list token) (s:noninitstate)
- (semantic_valuet:_),
- inhabited (parse_tree (last_symb_of_non_init_state s) wordt semantic_valuet) ->
-
- word_has_stack_semantics
- (wordq++wordt) (existT noninitstate_type s semantic_valuet::stackq).
-
-Lemma pop_invariant_converter:
- forall A symbols_to_pop symbols_popped,
- arrows_left (map symbol_semantic_type (rev_append symbols_to_pop symbols_popped)) A =
- arrows_left (map symbol_semantic_type symbols_popped)
- (arrows_right A (map symbol_semantic_type symbols_to_pop)).
-Proof.
-intros.
-unfold arrows_right, arrows_left.
-rewrite rev_append_rev, map_app, map_rev, fold_left_app.
-apply f_equal.
-rewrite <- fold_left_rev_right, rev_involutive.
-reflexivity.
-Qed.
-
-(** [pop] preserves the invariant **)
-Lemma pop_invariant:
- forall (symbols_to_pop symbols_popped:list symbol)
- (stack_cur:stack)
- (A:Type)
- (action:arrows_left (map symbol_semantic_type (rev_append symbols_to_pop symbols_popped)) A),
- forall word_stack word_popped,
- forall sem_popped,
- word_has_stack_semantics word_stack stack_cur ->
- inhabited (parse_tree_list symbols_popped word_popped sem_popped) ->
- let action' := eq_rect _ (fun x=>x) action _ (pop_invariant_converter _ _ _) in
- match pop symbols_to_pop stack_cur (uncurry action' sem_popped) with
- | OK (stack_new, sem) =>
- exists word1res word2res sem_full,
- (word_stack = word1res ++ word2res)%list /\
- word_has_stack_semantics word1res stack_new /\
- sem = uncurry action sem_full /\
- inhabited (
- parse_tree_list (rev_append symbols_to_pop symbols_popped) (word2res++word_popped) sem_full)
- | Err => True
- end.
-Proof.
-induction symbols_to_pop; intros; unfold pop; fold pop.
-exists word_stack, ([]:list token), sem_popped; intuition.
-f_equal.
-apply JMeq_eq, JMeq_eqrect with (P:=(fun x => x)).
-destruct stack_cur as [|[]]; eauto.
-destruct (compare_eqdec (last_symb_of_non_init_state x) a); eauto.
-destruct e; simpl.
-dependent destruction H.
-destruct H0, H1. apply (Cons_ptl X), inhabits in X0.
-specialize (IHsymbols_to_pop _ _ _ action0 _ _ _ H X0).
-match goal with
- IHsymbols_to_pop:match ?p1 with Err => _ | OK _ => _ end |- match ?p2 with Err => _ | OK _ => _ end =>
- replace p2 with p1; [destruct p1 as [|[]]|]; intuition
-end.
-destruct IHsymbols_to_pop as [word1res [word2res [sem_full []]]]; intuition; subst.
-exists word1res.
-eexists.
-exists sem_full.
-intuition.
-rewrite <- app_assoc; assumption.
-simpl; f_equal; f_equal.
-apply JMeq_eq.
-etransitivity.
-apply JMeq_eqrect with (P:=(fun x => x)).
-symmetry.
-apply JMeq_eqrect with (P:=(fun x => x)).
-Qed.
-
-(** [reduce_step] preserves the invariant **)
-Lemma reduce_step_invariant (stack:stack) (prod:production):
- forall word buffer, word_has_stack_semantics word stack ->
- match reduce_step init stack prod buffer with
- | OK (Accept_sr sem buffer_new) =>
- buffer = buffer_new /\
- inhabited (parse_tree (NT (start_nt init)) word sem)
- | OK (Progress_sr stack_new buffer_new) =>
- buffer = buffer_new /\
- word_has_stack_semantics word stack_new
- | Err | OK Fail_sr => True
- end.
-Proof with eauto.
-intros.
-unfold reduce_step.
-pose proof (pop_invariant (prod_rhs_rev prod) [] stack (symbol_semantic_type (NT (prod_lhs prod)))).
-revert H0.
-generalize (pop_invariant_converter (symbol_semantic_type (NT (prod_lhs prod))) (prod_rhs_rev prod) []).
-rewrite <- rev_alt.
-intros.
-specialize (H0 (prod_action prod) _ [] () H (inhabits Nil_ptl)).
-match goal with H0:let action' := ?a in _ |- _ => replace a with (prod_action' prod) in H0 end.
-simpl in H0.
-destruct (pop (prod_rhs_rev prod) stack (prod_action' prod)) as [|[]]; intuition.
-destruct H0 as [word1res [word2res [sem_full]]]; intuition.
-destruct H4; apply Non_terminal_pt, inhabits in X.
-match goal with X:inhabited (parse_tree _ _ ?u) |- _ => replace u with s0 in X end.
-clear sem_full H2.
-simpl; destruct (goto_table (state_of_stack init s) (prod_lhs prod)) as [[]|]; intuition; subst.
-rewrite app_nil_r in X; revert s0 X; rewrite e0; intro; simpl.
-constructor...
-destruct s; intuition.
-destruct (compare_eqdec (prod_lhs prod) (start_nt init)); intuition.
-rewrite app_nil_r in X.
-rewrite <- e0.
-inversion H0.
-destruct X; constructor...
-apply JMeq_eq.
-etransitivity.
-apply JMeq_eqrect with (P:=(fun x => x)).
-symmetry.
-apply JMeq_eqrect with (P:=(fun x => x)).
-Qed.
-
-(** [step] preserves the invariant **)
-Lemma step_invariant (stack:stack) (buffer:Stream token):
- forall buffer_tmp,
- (exists word_old,
- buffer = word_old ++ buffer_tmp /\
- word_has_stack_semantics word_old stack) ->
- match step init stack buffer_tmp with
- | OK (Accept_sr sem buffer_new) =>
- exists word_new,
- buffer = word_new ++ buffer_new /\
- inhabited (parse_tree (NT (start_nt init)) word_new sem)
- | OK (Progress_sr stack_new buffer_new) =>
- exists word_new,
- buffer = word_new ++ buffer_new /\
- word_has_stack_semantics word_new stack_new
- | Err | OK Fail_sr => True
- end.
-Proof with eauto.
-intros.
-destruct H, H.
-unfold step.
-destruct (action_table (state_of_stack init stack)).
-pose proof (reduce_step_invariant stack p x buffer_tmp).
-destruct (reduce_step init stack p buffer_tmp) as [|[]]; intuition; subst...
-destruct buffer_tmp.
-unfold Streams.hd.
-destruct t.
-destruct (l x0); intuition.
-exists (x ++ [existT (fun t => symbol_semantic_type (T t)) x0 s])%list.
-split.
-now rewrite <- app_str_app_assoc; intuition.
-apply Cons_stack_whss; intuition.
-destruct e; simpl.
-now exact (inhabits (Terminal_pt _ _)).
-match goal with |- (match reduce_step init stack p ?buff with Err => _ | OK _ => _ end) =>
- pose proof (reduce_step_invariant stack p x buff);
- destruct (reduce_step init stack p buff) as [|[]]; intuition; subst
-end...
-Qed.
-
-(** The interpreter is correct : if it returns a semantic value, then the input
- word has this semantic value.
-**)
-Theorem parse_correct buffer n_steps:
- match parse init buffer n_steps with
- | OK (Parsed_pr sem buffer_new) =>
- exists word_new,
- buffer = word_new ++ buffer_new /\
- inhabited (parse_tree (NT (start_nt init)) word_new sem)
- | _ => True
- end.
-Proof.
-unfold parse.
-assert (exists w, buffer = w ++ buffer /\ word_has_stack_semantics w []).
-exists ([]:list token); intuition.
-now apply Nil_stack_whss.
-revert H.
-generalize ([]:stack), buffer at 2 3.
-induction n_steps; simpl; intuition.
-pose proof (step_invariant _ _ _ H).
-destruct (step init s buffer0); simpl; intuition.
-destruct s0; intuition.
-apply IHn_steps; intuition.
-Qed.
-
-End Init.
-
-End Make.
diff --git a/cparser/MenhirLib/Interpreter_safe.v b/cparser/MenhirLib/Interpreter_safe.v
deleted file mode 100644
index a1aa35b8..00000000
--- a/cparser/MenhirLib/Interpreter_safe.v
+++ /dev/null
@@ -1,275 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Jacques-Henri Jourdan, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-Require Import Streams.
-Require Import Equality.
-Require Import List.
-Require Import Syntax.
-Require Import Alphabet.
-Require Grammar.
-Require Automaton.
-Require Validator_safe.
-Require Interpreter.
-
-Module Make(Import A:Automaton.T) (Import Inter:Interpreter.T A).
-Module Import Valid := Validator_safe.Make A.
-
-(** * A correct automaton does not crash **)
-
-Section Safety_proof.
-
-Hypothesis safe: safe.
-
-Proposition shift_head_symbs: shift_head_symbs.
-Proof. pose proof safe; unfold Valid.safe in H; intuition. Qed.
-Proposition goto_head_symbs: goto_head_symbs.
-Proof. pose proof safe; unfold Valid.safe in H; intuition. Qed.
-Proposition shift_past_state: shift_past_state.
-Proof. pose proof safe; unfold Valid.safe in H; intuition. Qed.
-Proposition goto_past_state: goto_past_state.
-Proof. pose proof safe; unfold Valid.safe in H; intuition. Qed.
-Proposition reduce_ok: reduce_ok.
-Proof. pose proof safe; unfold Valid.safe in H; intuition. Qed.
-
-(** We prove that a correct automaton won't crash : the interpreter will
- not return [Err] **)
-
-Variable init : initstate.
-
-(** The stack of states of an automaton stack **)
-Definition state_stack_of_stack (stack:stack) :=
- (List.map
- (fun cell:sigT noninitstate_type => singleton_state_pred (projT1 cell))
- stack ++ [singleton_state_pred init])%list.
-
-(** The stack of symbols of an automaton stack **)
-Definition symb_stack_of_stack (stack:stack) :=
- List.map
- (fun cell:sigT noninitstate_type => last_symb_of_non_init_state (projT1 cell))
- stack.
-
-(** The stack invariant : it basically states that the assumptions on the
- states are true. **)
-Inductive stack_invariant: stack -> Prop :=
- | stack_invariant_constr:
- forall stack,
- prefix (head_symbs_of_state (state_of_stack init stack))
- (symb_stack_of_stack stack) ->
- prefix_pred (head_states_of_state (state_of_stack init stack))
- (state_stack_of_stack stack) ->
- stack_invariant_next stack ->
- stack_invariant stack
-with stack_invariant_next: stack -> Prop :=
- | stack_invariant_next_nil:
- stack_invariant_next []
- | stack_invariant_next_cons:
- forall state_cur st stack_rec,
- stack_invariant stack_rec ->
- stack_invariant_next (existT _ state_cur st::stack_rec).
-
-(** [pop] conserves the stack invariant and does not crash
- under the assumption that we can pop at this place.
- Moreover, after a pop, the top state of the stack is allowed. **)
-Lemma pop_stack_invariant_conserved
- (symbols_to_pop:list symbol) (stack_cur:stack) A action:
- stack_invariant stack_cur ->
- prefix symbols_to_pop (head_symbs_of_state (state_of_stack init stack_cur)) ->
- match pop symbols_to_pop stack_cur (A:=A) action with
- | OK (stack_new, _) =>
- stack_invariant stack_new /\
- state_valid_after_pop
- (state_of_stack init stack_new) symbols_to_pop
- (head_states_of_state (state_of_stack init stack_cur))
- | Err => False
- end.
-Proof with eauto.
- intros.
- pose proof H.
- destruct H.
- revert H H0 H1 H2 H3.
- generalize (head_symbs_of_state (state_of_stack init stack0)).
- generalize (head_states_of_state (state_of_stack init stack0)).
- revert stack0 A action.
- induction symbols_to_pop; intros.
- - split...
- destruct l; constructor.
- inversion H2; subst.
- specialize (H7 (state_of_stack init stack0)).
- destruct (f2 (state_of_stack init stack0)) as [] eqn:? ...
- destruct stack0 as [|[]]; simpl in *.
- + inversion H6; subst.
- unfold singleton_state_pred in Heqb0.
- now rewrite compare_refl in Heqb0; discriminate.
- + inversion H6; subst.
- unfold singleton_state_pred in Heqb0.
- now rewrite compare_refl in Heqb0; discriminate.
- - destruct stack0 as [|[]]; unfold pop.
- + inversion H0; subst.
- now inversion H.
- + fold pop.
- destruct (compare_eqdec (last_symb_of_non_init_state x) a).
- * inversion H0; subst. clear H0.
- inversion H; subst. clear H.
- dependent destruction H3; simpl.
- assert (prefix_pred (List.tl l) (state_stack_of_stack stack0)).
- unfold tl; destruct l; [constructor | inversion H2]...
- pose proof H. destruct H3.
- specialize (IHsymbols_to_pop stack0 A (action0 n) _ _ H4 H7 H H0 H6).
- revert IHsymbols_to_pop.
- fold (noninitstate_type x); generalize (pop symbols_to_pop stack0 (action0 n)).
- destruct r as [|[]]; intuition...
- destruct l; constructor...
- * apply n0.
- inversion H0; subst.
- inversion H; subst...
-Qed.
-
-(** [prefix] is associative **)
-Lemma prefix_ass:
- forall (l1 l2 l3:list symbol), prefix l1 l2 -> prefix l2 l3 -> prefix l1 l3.
-Proof.
-induction l1; intros.
-constructor.
-inversion H; subst.
-inversion H0; subst.
-constructor; eauto.
-Qed.
-
-(** [prefix_pred] is associative **)
-Lemma prefix_pred_ass:
- forall (l1 l2 l3:list (state->bool)),
- prefix_pred l1 l2 -> prefix_pred l2 l3 -> prefix_pred l1 l3.
-Proof.
-induction l1; intros.
-constructor.
-inversion H; subst.
-inversion H0; subst.
-constructor; eauto.
-intro.
-specialize (H3 x).
-specialize (H4 x).
-destruct (f0 x); simpl in *; intuition.
-rewrite H4 in H3; intuition.
-Qed.
-
-(** If we have the right to reduce at this state, then the stack invariant
- is conserved by [reduce_step] and [reduce_step] does not crash. **)
-Lemma reduce_step_stack_invariant_conserved stack prod buff:
- stack_invariant stack ->
- valid_for_reduce (state_of_stack init stack) prod ->
- match reduce_step init stack prod buff with
- | OK (Fail_sr | Accept_sr _ _) => True
- | OK (Progress_sr stack_new _) => stack_invariant stack_new
- | Err => False
- end.
-Proof with eauto.
-unfold valid_for_reduce.
-intuition.
-unfold reduce_step.
-pose proof (pop_stack_invariant_conserved (prod_rhs_rev prod) stack _ (prod_action' prod)).
-destruct (pop (prod_rhs_rev prod) stack (prod_action' prod)) as [|[]].
-apply H0...
-destruct H0...
-pose proof (goto_head_symbs (state_of_stack init s) (prod_lhs prod)).
-pose proof (goto_past_state (state_of_stack init s) (prod_lhs prod)).
-unfold bind2.
-destruct H0.
-specialize (H2 _ H3)...
-destruct (goto_table (state_of_stack init stack0) (prod_lhs prod)) as [[]|].
-constructor.
-simpl.
-constructor.
-eapply prefix_ass...
-unfold state_stack_of_stack; simpl; constructor.
-intro; destruct (singleton_state_pred x x0); reflexivity.
-eapply prefix_pred_ass...
-constructor...
-constructor...
-destruct stack0 as [|[]]...
-destruct (compare_eqdec (prod_lhs prod) (start_nt init))...
-apply n, H2, eq_refl.
-apply H2, eq_refl.
-Qed.
-
-(** If the automaton is safe, then the stack invariant is conserved by
- [step] and [step] does not crash. **)
-Lemma step_stack_invariant_conserved (stack:stack) buffer:
- stack_invariant stack ->
- match step init stack buffer with
- | OK (Fail_sr | Accept_sr _ _) => True
- | OK (Progress_sr stack_new _) => stack_invariant stack_new
- | Err => False
- end.
-Proof with eauto.
-intro.
-unfold step.
-pose proof (shift_head_symbs (state_of_stack init stack)).
-pose proof (shift_past_state (state_of_stack init stack)).
-pose proof (reduce_ok (state_of_stack init stack)).
-destruct (action_table (state_of_stack init stack)).
-apply reduce_step_stack_invariant_conserved...
-destruct buffer as [[]]; simpl.
-specialize (H0 x); specialize (H1 x); specialize (H2 x).
-destruct (l x)...
-destruct H.
-constructor.
-unfold state_of_stack.
-constructor.
-eapply prefix_ass...
-unfold state_stack_of_stack; simpl; constructor.
-intro; destruct (singleton_state_pred s0 x0)...
-eapply prefix_pred_ass...
-constructor...
-constructor...
-apply reduce_step_stack_invariant_conserved...
-Qed.
-
-(** If the automaton is safe, then it does not crash **)
-Theorem parse_no_err buffer n_steps:
- parse init buffer n_steps <> Err.
-Proof with eauto.
-unfold parse.
-assert (stack_invariant []).
-constructor.
-constructor.
-unfold state_stack_of_stack; simpl; constructor.
-intro; destruct (singleton_state_pred init x)...
-constructor.
-constructor.
-revert H.
-generalize buffer ([]:stack).
-induction n_steps; simpl.
-now discriminate.
-intros.
-pose proof (step_stack_invariant_conserved s buffer0 H).
-destruct (step init s buffer0) as [|[]]; simpl...
-discriminate.
-discriminate.
-Qed.
-
-(** A version of [parse] that uses safeness in order to return a
- [parse_result], and not a [result parse_result] : we have proven that
- parsing does not return [Err]. **)
-Definition parse_with_safe (buffer:Stream token) (n_steps:nat):
- parse_result init.
-Proof with eauto.
-pose proof (parse_no_err buffer n_steps).
-destruct (parse init buffer n_steps)...
-congruence.
-Defined.
-
-End Safety_proof.
-
-End Make.
diff --git a/cparser/MenhirLib/Main.v b/cparser/MenhirLib/Main.v
deleted file mode 100644
index 1a17e988..00000000
--- a/cparser/MenhirLib/Main.v
+++ /dev/null
@@ -1,92 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Jacques-Henri Jourdan, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-Require Grammar.
-Require Automaton.
-Require Interpreter_safe.
-Require Interpreter_correct.
-Require Interpreter_complete.
-Require Import Syntax.
-
-Module Make(Export Aut:Automaton.T).
-Export Aut.Gram.
-Export Aut.GramDefs.
-
-Module Import Inter := Interpreter.Make Aut.
-Module Safe := Interpreter_safe.Make Aut Inter.
-Module Correct := Interpreter_correct.Make Aut Inter.
-Module Complete := Interpreter_complete.Make Aut Inter.
-
-Definition complete_validator:unit->bool := Complete.Valid.is_complete.
-Definition safe_validator:unit->bool := Safe.Valid.is_safe.
-Definition parse (safe:safe_validator ()=true) init n_steps buffer : parse_result init:=
- Safe.parse_with_safe (Safe.Valid.is_safe_correct safe) init buffer n_steps.
-
-(** Correction theorem. **)
-Theorem parse_correct
- (safe:safe_validator ()= true) init n_steps buffer:
- match parse safe init n_steps buffer with
- | Parsed_pr sem buffer_new =>
- exists word,
- buffer = word ++ buffer_new /\ inhabited (parse_tree (NT (start_nt init)) word sem)
- | _ => True
- end.
-Proof.
-unfold parse, Safe.parse_with_safe.
-pose proof (Correct.parse_correct init buffer n_steps).
-generalize (Safe.parse_no_err (Safe.Valid.is_safe_correct safe) init buffer n_steps).
-destruct (Inter.parse init buffer n_steps); intros.
-now destruct (n (eq_refl _)).
-now destruct p; trivial.
-Qed.
-
-(** Completeness theorem. **)
-Theorem parse_complete
- (safe:safe_validator () = true) init n_steps word buffer_end sem:
- complete_validator () = true ->
- forall tree:parse_tree (NT (start_nt init)) word sem,
- match parse safe init n_steps (word ++ buffer_end) with
- | Fail_pr => False
- | Parsed_pr sem_res buffer_end_res =>
- sem_res = sem /\ buffer_end_res = buffer_end /\ pt_size tree <= n_steps
- | Timeout_pr => n_steps < pt_size tree
- end.
-Proof.
-intros.
-unfold parse, Safe.parse_with_safe.
-pose proof (Complete.parse_complete (Complete.Valid.is_complete_correct H) init _ buffer_end _ tree n_steps).
-generalize (Safe.parse_no_err (Safe.Valid.is_safe_correct safe) init (word ++ buffer_end) n_steps).
-destruct (Inter.parse init (word ++ buffer_end) n_steps); intros.
-now destruct (n eq_refl).
-now exact H0.
-Qed.
-
-(** Unambiguity theorem. **)
-Theorem unambiguity:
- safe_validator () = true -> complete_validator () = true -> inhabited token ->
- forall init word,
- forall sem1 (tree1:parse_tree (NT (start_nt init)) word sem1),
- forall sem2 (tree2:parse_tree (NT (start_nt init)) word sem2),
- sem1 = sem2.
-Proof.
-intros.
-destruct H1.
-pose proof (parse_complete H init (pt_size tree1) word (Streams.const X) sem1) H0 tree1.
-pose proof (parse_complete H init (pt_size tree1) word (Streams.const X) sem2) H0 tree2.
-destruct (parse H init (pt_size tree1) (word ++ Streams.const X)); intuition.
-rewrite <- H3, H1; reflexivity.
-Qed.
-
-End Make.
diff --git a/cparser/MenhirLib/Tuples.v b/cparser/MenhirLib/Tuples.v
deleted file mode 100644
index 3fd2ec03..00000000
--- a/cparser/MenhirLib/Tuples.v
+++ /dev/null
@@ -1,49 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Jacques-Henri Jourdan, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-Require Import List.
-Require Import Coq.Program.Syntax.
-Require Import Equality.
-
-(** A curryfied function with multiple parameters **)
-Definition arrows_left: list Type -> Type -> Type :=
- fold_left (fun A B => B -> A).
-
-(** A curryfied function with multiple parameters **)
-Definition arrows_right: Type -> list Type -> Type :=
- fold_right (fun A B => A -> B).
-
-(** A tuple is a heterogeneous list. For convenience, we use pairs. **)
-Fixpoint tuple (types : list Type) : Type :=
- match types with
- | nil => unit
- | t::q => prod t (tuple q)
- end.
-
-Fixpoint uncurry {args:list Type} {res:Type}:
- arrows_left args res -> tuple args -> res :=
- match args return forall res, arrows_left args res -> tuple args -> res with
- | [] => fun _ f _ => f
- | t::q => fun res f p => let (d, t) := p in
- (@uncurry q _ f t) d
- end res.
-
-Lemma JMeq_eqrect:
- forall (U:Type) (a b:U) (P:U -> Type) (x:P a) (e:a=b),
- eq_rect a P x b e ~= x.
-Proof.
-destruct e.
-reflexivity.
-Qed.
diff --git a/cparser/MenhirLib/Validator_complete.v b/cparser/MenhirLib/Validator_complete.v
deleted file mode 100644
index a9823278..00000000
--- a/cparser/MenhirLib/Validator_complete.v
+++ /dev/null
@@ -1,542 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Jacques-Henri Jourdan, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-Require Automaton.
-Require Import Alphabet.
-Require Import List.
-Require Import Syntax.
-
-Module Make(Import A:Automaton.T).
-
-(** We instantiate some sets/map. **)
-Module TerminalComparableM <: ComparableM.
- Definition t := terminal.
- Instance tComparable : Comparable t := _.
-End TerminalComparableM.
-Module TerminalOrderedType := OrderedType_from_ComparableM TerminalComparableM.
-Module StateProdPosComparableM <: ComparableM.
- Definition t := (state*production*nat)%type.
- Instance tComparable : Comparable t := _.
-End StateProdPosComparableM.
-Module StateProdPosOrderedType :=
- OrderedType_from_ComparableM StateProdPosComparableM.
-
-Module TerminalSet := FSetAVL.Make TerminalOrderedType.
-Module StateProdPosMap := FMapAVL.Make StateProdPosOrderedType.
-
-(** Nullable predicate for symbols and list of symbols. **)
-Definition nullable_symb (symbol:symbol) :=
- match symbol with
- | NT nt => nullable_nterm nt
- | _ => false
- end.
-
-Definition nullable_word (word:list symbol) :=
- forallb nullable_symb word.
-
-(** First predicate for non terminal, symbols and list of symbols, given as FSets. **)
-Definition first_nterm_set (nterm:nonterminal) :=
- fold_left (fun acc t => TerminalSet.add t acc)
- (first_nterm nterm) TerminalSet.empty.
-
-Definition first_symb_set (symbol:symbol) :=
- match symbol with
- | NT nt => first_nterm_set nt
- | T t => TerminalSet.singleton t
- end.
-
-Fixpoint first_word_set (word:list symbol) :=
- match word with
- | [] => TerminalSet.empty
- | t::q =>
- if nullable_symb t then
- TerminalSet.union (first_symb_set t) (first_word_set q)
- else
- first_symb_set t
- end.
-
-(** Small helper for finding the part of an item that is after the dot. **)
-Definition future_of_prod prod dot_pos : list symbol :=
- (fix loop n lst :=
- match n with
- | O => lst
- | S x => match loop x lst with [] => [] | _::q => q end
- end)
- dot_pos (rev' (prod_rhs_rev prod)).
-
-(** We build a fast map to store all the items of all the states. **)
-Definition items_map (_:unit): StateProdPosMap.t TerminalSet.t :=
- fold_left (fun acc state =>
- fold_left (fun acc item =>
- let key := (state, prod_item item, dot_pos_item item) in
- let data := fold_left (fun acc t => TerminalSet.add t acc)
- (lookaheads_item item) TerminalSet.empty
- in
- let old :=
- match StateProdPosMap.find key acc with
- | Some x => x | None => TerminalSet.empty
- end
- in
- StateProdPosMap.add key (TerminalSet.union data old) acc
- ) (items_of_state state) acc
- ) all_list (StateProdPosMap.empty TerminalSet.t).
-
-(** Accessor. **)
-Definition find_items_map items_map state prod dot_pos : TerminalSet.t :=
- match StateProdPosMap.find (state, prod, dot_pos) items_map with
- | None => TerminalSet.empty
- | Some x => x
- end.
-
-Definition state_has_future state prod (fut:list symbol) (lookahead:terminal) :=
- exists dot_pos:nat,
- fut = future_of_prod prod dot_pos /\
- TerminalSet.In lookahead (find_items_map (items_map ()) state prod dot_pos).
-
-(** Iterator over items. **)
-Definition forallb_items items_map (P:state -> production -> nat -> TerminalSet.t -> bool): bool:=
- StateProdPosMap.fold (fun key set acc =>
- match key with (st, p, pos) => (acc && P st p pos set)%bool end
- ) items_map true.
-
-Lemma forallb_items_spec :
- forall p, forallb_items (items_map ()) p = true ->
- forall st prod fut lookahead, state_has_future st prod fut lookahead ->
- forall P:state -> production -> list symbol -> terminal -> Prop,
- (forall st prod pos set lookahead,
- TerminalSet.In lookahead set -> p st prod pos set = true ->
- P st prod (future_of_prod prod pos) lookahead) ->
- P st prod fut lookahead.
-Proof.
-intros.
-unfold forallb_items in H.
-rewrite StateProdPosMap.fold_1 in H.
-destruct H0; destruct H0.
-specialize (H1 st prod x _ _ H2).
-destruct H0.
-apply H1.
-unfold find_items_map in *.
-pose proof (@StateProdPosMap.find_2 _ (items_map ()) (st, prod, x)).
-destruct (StateProdPosMap.find (st, prod, x) (items_map ())); [ |destruct (TerminalSet.empty_1 H2)].
-specialize (H0 _ (eq_refl _)).
-pose proof (StateProdPosMap.elements_1 H0).
-revert H.
-generalize true at 1.
-induction H3.
-destruct H.
-destruct y.
-simpl in H3; destruct H3.
-pose proof (compare_eq (st, prod, x) k H).
-destruct H3.
-simpl.
-generalize (p st prod x t).
-induction l; simpl; intros.
-rewrite Bool.andb_true_iff in H3.
-intuition.
-destruct a; destruct k; destruct p0.
-simpl in H3.
-replace (b0 && b && p s p0 n t0)%bool with (b0 && p s p0 n t0 && b)%bool in H3.
-apply (IHl _ _ H3).
-destruct b, b0, (p s p0 n t0); reflexivity.
-intro.
-apply IHInA.
-Qed.
-
-(** * Validation for completeness **)
-
-(** The nullable predicate is a fixpoint : it is correct. **)
-Definition nullable_stable:=
- forall p:production,
- nullable_word (rev (prod_rhs_rev p)) = true ->
- nullable_nterm (prod_lhs p) = true.
-
-Definition is_nullable_stable (_:unit) :=
- forallb (fun p:production =>
- implb (nullable_word (rev' (prod_rhs_rev p))) (nullable_nterm (prod_lhs p)))
- all_list.
-
-Property is_nullable_stable_correct :
- is_nullable_stable () = true -> nullable_stable.
-Proof.
-unfold is_nullable_stable, nullable_stable.
-intros.
-rewrite forallb_forall in H.
-specialize (H p (all_list_forall p)).
-unfold rev' in H; rewrite <- rev_alt in H.
-rewrite H0 in H; intuition.
-Qed.
-
-(** The first predicate is a fixpoint : it is correct. **)
-Definition first_stable:=
- forall (p:production),
- TerminalSet.Subset (first_word_set (rev (prod_rhs_rev p)))
- (first_nterm_set (prod_lhs p)).
-
-Definition is_first_stable (_:unit) :=
- forallb (fun p:production =>
- TerminalSet.subset (first_word_set (rev' (prod_rhs_rev p)))
- (first_nterm_set (prod_lhs p)))
- all_list.
-
-Property is_first_stable_correct :
- is_first_stable () = true -> first_stable.
-Proof.
-unfold is_first_stable, first_stable.
-intros.
-rewrite forallb_forall in H.
-specialize (H p (all_list_forall p)).
-unfold rev' in H; rewrite <- rev_alt in H.
-apply TerminalSet.subset_2; intuition.
-Qed.
-
-(** The initial state has all the S=>.u items, where S is the start non-terminal **)
-Definition start_future :=
- forall (init:initstate) (t:terminal) (p:production),
- prod_lhs p = start_nt init ->
- state_has_future init p (rev (prod_rhs_rev p)) t.
-
-Definition is_start_future items_map :=
- forallb (fun init =>
- forallb (fun prod =>
- if compare_eqb (prod_lhs prod) (start_nt init) then
- let lookaheads := find_items_map items_map init prod 0 in
- forallb (fun t => TerminalSet.mem t lookaheads) all_list
- else
- true) all_list) all_list.
-
-Property is_start_future_correct :
- is_start_future (items_map ()) = true -> start_future.
-Proof.
-unfold is_start_future, start_future.
-intros.
-rewrite forallb_forall in H.
-specialize (H init (all_list_forall _)).
-rewrite forallb_forall in H.
-specialize (H p (all_list_forall _)).
-rewrite <- compare_eqb_iff in H0.
-rewrite H0 in H.
-rewrite forallb_forall in H.
-specialize (H t (all_list_forall _)).
-exists 0.
-split.
-apply rev_alt.
-apply TerminalSet.mem_2; eauto.
-Qed.
-
-(** If a state contains an item of the form A->_.av[[b]], where a is a
- terminal, then reading an a does a [Shift_act], to a state containing
- an item of the form A->_.v[[b]]. **)
-Definition terminal_shift :=
- forall (s1:state) prod fut lookahead,
- state_has_future s1 prod fut lookahead ->
- match fut with
- | T t::q =>
- match action_table s1 with
- | Lookahead_act awp =>
- match awp t with
- | Shift_act s2 _ =>
- state_has_future s2 prod q lookahead
- | _ => False
- end
- | _ => False
- end
- | _ => True
- end.
-
-Definition is_terminal_shift items_map :=
- forallb_items items_map (fun s1 prod pos lset =>
- match future_of_prod prod pos with
- | T t::_ =>
- match action_table s1 with
- | Lookahead_act awp =>
- match awp t with
- | Shift_act s2 _ =>
- TerminalSet.subset lset (find_items_map items_map s2 prod (S pos))
- | _ => false
- end
- | _ => false
- end
- | _ => true
- end).
-
-Property is_terminal_shift_correct :
- is_terminal_shift (items_map ()) = true -> terminal_shift.
-Proof.
-unfold is_terminal_shift, terminal_shift.
-intros.
-apply (forallb_items_spec _ H _ _ _ _ H0 (fun _ _ fut look => _)).
-intros.
-destruct (future_of_prod prod0 pos) as [|[]] eqn:?; intuition.
-destruct (action_table st); intuition.
-destruct (l0 t); intuition.
-exists (S pos).
-split.
-unfold future_of_prod in *.
-rewrite Heql; reflexivity.
-apply (TerminalSet.subset_2 H2); intuition.
-Qed.
-
-(** If a state contains an item of the form A->_.[[a]], then either we do a
- [Default_reduce_act] of the corresponding production, either a is a
- terminal (ie. there is a lookahead terminal), and reading a does a
- [Reduce_act] of the corresponding production. **)
-Definition end_reduce :=
- forall (s:state) prod fut lookahead,
- state_has_future s prod fut lookahead ->
- fut = [] ->
- match action_table s with
- | Default_reduce_act p => p = prod
- | Lookahead_act awt =>
- match awt lookahead with
- | Reduce_act p => p = prod
- | _ => False
- end
- end.
-
-Definition is_end_reduce items_map :=
- forallb_items items_map (fun s prod pos lset =>
- match future_of_prod prod pos with
- | [] =>
- match action_table s with
- | Default_reduce_act p => compare_eqb p prod
- | Lookahead_act awt =>
- TerminalSet.fold (fun lookahead acc =>
- match awt lookahead with
- | Reduce_act p => (acc && compare_eqb p prod)%bool
- | _ => false
- end) lset true
- end
- | _ => true
- end).
-
-Property is_end_reduce_correct :
- is_end_reduce (items_map ()) = true -> end_reduce.
-Proof.
-unfold is_end_reduce, end_reduce.
-intros.
-revert H1.
-apply (forallb_items_spec _ H _ _ _ _ H0 (fun st prod fut look => _ ->
- match action_table st with
- | Default_reduce_act p => p = prod
- | _ => _
- end)).
-intros.
-rewrite H3 in H2.
-destruct (action_table st); intuition.
-apply compare_eqb_iff; intuition.
-rewrite TerminalSet.fold_1 in H2.
-revert H2.
-generalize true at 1.
-pose proof (TerminalSet.elements_1 H1).
-induction H2.
-pose proof (compare_eq _ _ H2).
-destruct H4.
-simpl.
-assert (fold_left
- (fun (a : bool) (e : TerminalSet.elt) =>
- match l e with
- | Shift_act _ _ => false
- | Reduce_act p => (a && compare_eqb p prod0)%bool
- | Fail_act => false
- end) l0 false = true -> False).
-induction l0; intuition.
-apply IHl0.
-simpl in H4.
-destruct (l a); intuition.
-destruct (l lookahead0); intuition.
-apply compare_eqb_iff.
-destruct (compare_eqb p prod0); intuition.
-destruct b; intuition.
-simpl; intros.
-eapply IHInA; eauto.
-Qed.
-
-(** If a state contains an item of the form A->_.Bv[[b]], where B is a
- non terminal, then the goto table says we have to go to a state containing
- an item of the form A->_.v[[b]]. **)
-Definition non_terminal_goto :=
- forall (s1:state) prod fut lookahead,
- state_has_future s1 prod fut lookahead ->
- match fut with
- | NT nt::q =>
- match goto_table s1 nt with
- | Some (exist _ s2 _) =>
- state_has_future s2 prod q lookahead
- | None =>
- forall prod fut lookahead,
- state_has_future s1 prod fut lookahead ->
- match fut with
- | NT nt'::_ => nt <> nt'
- | _ => True
- end
- end
- | _ => True
- end.
-
-Definition is_non_terminal_goto items_map :=
- forallb_items items_map (fun s1 prod pos lset =>
- match future_of_prod prod pos with
- | NT nt::_ =>
- match goto_table s1 nt with
- | Some (exist _ s2 _) =>
- TerminalSet.subset lset (find_items_map items_map s2 prod (S pos))
- | None => forallb_items items_map (fun s1' prod' pos' _ =>
- (implb (compare_eqb s1 s1')
- match future_of_prod prod' pos' with
- | NT nt' :: _ => negb (compare_eqb nt nt')
- | _ => true
- end)%bool)
- end
- | _ => true
- end).
-
-Property is_non_terminal_goto_correct :
- is_non_terminal_goto (items_map ()) = true -> non_terminal_goto.
-Proof.
-unfold is_non_terminal_goto, non_terminal_goto.
-intros.
-apply (forallb_items_spec _ H _ _ _ _ H0 (fun st prod fut look =>
- match fut with
- | NT nt :: q =>
- match goto_table st nt with
- | Some _ => _
- | None =>
- forall p f l, state_has_future st p f l -> (_:Prop)
- end
- | _ => _
- end)).
-intros.
-destruct (future_of_prod prod0 pos) as [|[]] eqn:?; intuition.
-destruct (goto_table st n) as [[]|].
-exists (S pos).
-split.
-unfold future_of_prod in *.
-rewrite Heql; reflexivity.
-apply (TerminalSet.subset_2 H2); intuition.
-intros.
-remember st in H2; revert Heqs.
-apply (forallb_items_spec _ H2 _ _ _ _ H3 (fun st' prod fut look => s = st' -> match fut return Prop with [] => _ | _ => _ end)); intros.
-rewrite <- compare_eqb_iff in H6; rewrite H6 in H5.
-destruct (future_of_prod prod1 pos0) as [|[]]; intuition.
-rewrite <- compare_eqb_iff in H7; rewrite H7 in H5.
-discriminate.
-Qed.
-
-Definition start_goto :=
- forall (init:initstate), goto_table init (start_nt init) = None.
-
-Definition is_start_goto (_:unit) :=
- forallb (fun (init:initstate) =>
- match goto_table init (start_nt init) with
- | Some _ => false
- | None => true
- end) all_list.
-
-Definition is_start_goto_correct:
- is_start_goto () = true -> start_goto.
-Proof.
-unfold is_start_goto, start_goto.
-rewrite forallb_forall.
-intros.
-specialize (H init (all_list_forall _)).
-destruct (goto_table init (start_nt init)); congruence.
-Qed.
-
-(** Closure property of item sets : if a state contains an item of the form
- A->_.Bv[[b]], then for each production B->u and each terminal a of
- first(vb), the state contains an item of the form B->_.u[[a]] **)
-Definition non_terminal_closed :=
- forall (s1:state) prod fut lookahead,
- state_has_future s1 prod fut lookahead ->
- match fut with
- | NT nt::q =>
- forall (p:production) (lookahead2:terminal),
- prod_lhs p = nt ->
- TerminalSet.In lookahead2 (first_word_set q) \/
- lookahead2 = lookahead /\ nullable_word q = true ->
- state_has_future s1 p (rev (prod_rhs_rev p)) lookahead2
- | _ => True
- end.
-
-Definition is_non_terminal_closed items_map :=
- forallb_items items_map (fun s1 prod pos lset =>
- match future_of_prod prod pos with
- | NT nt::q =>
- forallb (fun p =>
- if compare_eqb (prod_lhs p) nt then
- let lookaheads := find_items_map items_map s1 p 0 in
- (implb (nullable_word q) (TerminalSet.subset lset lookaheads)) &&
- TerminalSet.subset (first_word_set q) lookaheads
- else true
- )%bool all_list
- | _ => true
- end).
-
-Property is_non_terminal_closed_correct:
- is_non_terminal_closed (items_map ()) = true -> non_terminal_closed.
-Proof.
-unfold is_non_terminal_closed, non_terminal_closed.
-intros.
-apply (forallb_items_spec _ H _ _ _ _ H0 (fun st prod fut look =>
- match fut with
- | NT nt :: q => forall p l, _ -> _ -> state_has_future st _ _ _
- | _ => _
- end)).
-intros.
-destruct (future_of_prod prod0 pos); intuition.
-destruct s; eauto; intros.
-rewrite forallb_forall in H2.
-specialize (H2 p (all_list_forall p)).
-rewrite <- compare_eqb_iff in H3.
-rewrite H3 in H2.
-rewrite Bool.andb_true_iff in H2.
-destruct H2.
-exists 0.
-split.
-apply rev_alt.
-destruct H4 as [|[]]; subst.
-apply (TerminalSet.subset_2 H5); intuition.
-rewrite H6 in H2.
-apply (TerminalSet.subset_2 H2); intuition.
-Qed.
-
-(** The automaton is complete **)
-Definition complete :=
- nullable_stable /\ first_stable /\ start_future /\ terminal_shift
- /\ end_reduce /\ non_terminal_goto /\ start_goto /\ non_terminal_closed.
-
-Definition is_complete (_:unit) :=
- let items_map := items_map () in
- (is_nullable_stable () && is_first_stable () && is_start_future items_map &&
- is_terminal_shift items_map && is_end_reduce items_map && is_start_goto () &&
- is_non_terminal_goto items_map && is_non_terminal_closed items_map)%bool.
-
-Property is_complete_correct:
- is_complete () = true -> complete.
-Proof.
-unfold is_complete, complete.
-repeat rewrite Bool.andb_true_iff.
-intuition.
-apply is_nullable_stable_correct; assumption.
-apply is_first_stable_correct; assumption.
-apply is_start_future_correct; assumption.
-apply is_terminal_shift_correct; assumption.
-apply is_end_reduce_correct; assumption.
-apply is_non_terminal_goto_correct; assumption.
-apply is_start_goto_correct; assumption.
-apply is_non_terminal_closed_correct; assumption.
-Qed.
-
-End Make.
diff --git a/cparser/MenhirLib/Validator_safe.v b/cparser/MenhirLib/Validator_safe.v
deleted file mode 100644
index 183d661b..00000000
--- a/cparser/MenhirLib/Validator_safe.v
+++ /dev/null
@@ -1,414 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Jacques-Henri Jourdan, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-Require Automaton.
-Require Import Alphabet.
-Require Import List.
-Require Import Syntax.
-
-Module Make(Import A:Automaton.T).
-
-(** The singleton predicate for states **)
-Definition singleton_state_pred (state:state) :=
- (fun state' => match compare state state' with Eq => true |_ => false end).
-
-(** [past_state_of_non_init_state], extended for all states. **)
-Definition past_state_of_state (state:state) :=
- match state with
- | Init _ => []
- | Ninit nis => past_state_of_non_init_state nis
- end.
-
-(** Concatenations of last and past **)
-Definition head_symbs_of_state (state:state) :=
- match state with
- | Init _ => []
- | Ninit s =>
- last_symb_of_non_init_state s::past_symb_of_non_init_state s
- end.
-Definition head_states_of_state (state:state) :=
- singleton_state_pred state::past_state_of_state state.
-
-(** * Validation for correctness **)
-
-(** Prefix predicate between two lists of symbols. **)
-Inductive prefix: list symbol -> list symbol -> Prop :=
- | prefix_nil: forall l, prefix [] l
- | prefix_cons: forall l1 l2 x, prefix l1 l2 -> prefix (x::l1) (x::l2).
-
-Fixpoint is_prefix (l1 l2:list symbol):=
- match l1, l2 with
- | [], _ => true
- | t1::q1, t2::q2 => (compare_eqb t1 t2 && is_prefix q1 q2)%bool
- | _::_, [] => false
- end.
-
-Property is_prefix_correct (l1 l2:list symbol):
- is_prefix l1 l2 = true -> prefix l1 l2.
-Proof.
-revert l2.
-induction l1; intros.
-apply prefix_nil.
-unfold is_prefix in H.
-destruct l2; intuition; try discriminate.
-rewrite Bool.andb_true_iff in H.
-destruct H.
-rewrite compare_eqb_iff in H.
-destruct H.
-apply prefix_cons.
-apply IHl1; intuition.
-Qed.
-
-(** If we shift, then the known top symbols of the destination state is
- a prefix of the known top symbols of the source state, with the new
- symbol added. **)
-Definition shift_head_symbs :=
- forall s,
- match action_table s with
- | Lookahead_act awp =>
- forall t, match awp t with
- | Shift_act s2 _ =>
- prefix (past_symb_of_non_init_state s2) (head_symbs_of_state s)
- | _ => True
- end
- | _ => True
- end.
-
-Definition is_shift_head_symbs (_:unit) :=
- forallb (fun s:state =>
- match action_table s with
- | Lookahead_act awp =>
- forallb (fun t =>
- match awp t with
- | Shift_act s2 _ =>
- is_prefix (past_symb_of_non_init_state s2) (head_symbs_of_state s)
- | _ => true
- end)
- all_list
- | _ => true
- end)
- all_list.
-
-Property is_shift_head_symbs_correct:
- is_shift_head_symbs () = true -> shift_head_symbs.
-Proof.
-unfold is_shift_head_symbs, shift_head_symbs.
-intros.
-rewrite forallb_forall in H.
-specialize (H s (all_list_forall s)).
-destruct (action_table s); intuition.
-rewrite forallb_forall in H.
-specialize (H t (all_list_forall t)).
-destruct (l t); intuition.
-apply is_prefix_correct; intuition.
-Qed.
-
-(** When a goto happens, then the known top symbols of the destination state
- is a prefix of the known top symbols of the source state, with the new
- symbol added. **)
-Definition goto_head_symbs :=
- forall s nt,
- match goto_table s nt with
- | Some (exist _ s2 _) =>
- prefix (past_symb_of_non_init_state s2) (head_symbs_of_state s)
- | None => True
- end.
-
-Definition is_goto_head_symbs (_:unit) :=
- forallb (fun s:state =>
- forallb (fun nt =>
- match goto_table s nt with
- | Some (exist _ s2 _) =>
- is_prefix (past_symb_of_non_init_state s2) (head_symbs_of_state s)
- | None => true
- end)
- all_list)
- all_list.
-
-Property is_goto_head_symbs_correct:
- is_goto_head_symbs () = true -> goto_head_symbs.
-Proof.
-unfold is_goto_head_symbs, goto_head_symbs.
-intros.
-rewrite forallb_forall in H.
-specialize (H s (all_list_forall s)).
-rewrite forallb_forall in H.
-specialize (H nt (all_list_forall nt)).
-destruct (goto_table s nt); intuition.
-destruct s0.
-apply is_prefix_correct; intuition.
-Qed.
-
-(** We have to say the same kind of checks for the assumptions about the
- states stack. However, theses assumptions are predicates. So we define
- a notion of "prefix" over predicates lists, that means, basically, that
- an assumption entails another **)
-Inductive prefix_pred: list (state->bool) -> list (state->bool) -> Prop :=
- | prefix_pred_nil: forall l, prefix_pred [] l
- | prefix_pred_cons: forall l1 l2 f1 f2,
- (forall x, implb (f2 x) (f1 x) = true) ->
- prefix_pred l1 l2 -> prefix_pred (f1::l1) (f2::l2).
-
-Fixpoint is_prefix_pred (l1 l2:list (state->bool)) :=
- match l1, l2 with
- | [], _ => true
- | f1::q1, f2::q2 =>
- (forallb (fun x => implb (f2 x) (f1 x)) all_list
- && is_prefix_pred q1 q2)%bool
- | _::_, [] => false
- end.
-
-Property is_prefix_pred_correct (l1 l2:list (state->bool)) :
- is_prefix_pred l1 l2 = true -> prefix_pred l1 l2.
-Proof.
-revert l2.
-induction l1.
-intros.
-apply prefix_pred_nil.
-intros.
-destruct l2; intuition; try discriminate.
-unfold is_prefix_pred in H.
-rewrite Bool.andb_true_iff in H.
-rewrite forallb_forall in H.
-apply prefix_pred_cons; intuition.
-apply H0.
-apply all_list_forall.
-Qed.
-
-(** The assumptions about state stack is conserved when we shift **)
-Definition shift_past_state :=
- forall s,
- match action_table s with
- | Lookahead_act awp =>
- forall t, match awp t with
- | Shift_act s2 _ =>
- prefix_pred (past_state_of_non_init_state s2)
- (head_states_of_state s)
- | _ => True
- end
- | _ => True
- end.
-
-Definition is_shift_past_state (_:unit) :=
- forallb (fun s:state =>
- match action_table s with
- | Lookahead_act awp =>
- forallb (fun t =>
- match awp t with
- | Shift_act s2 _ =>
- is_prefix_pred
- (past_state_of_non_init_state s2) (head_states_of_state s)
- | _ => true
- end)
- all_list
- | _ => true
- end)
- all_list.
-
-Property is_shift_past_state_correct:
- is_shift_past_state () = true -> shift_past_state.
-Proof.
-unfold is_shift_past_state, shift_past_state.
-intros.
-rewrite forallb_forall in H.
-specialize (H s (all_list_forall s)).
-destruct (action_table s); intuition.
-rewrite forallb_forall in H.
-specialize (H t (all_list_forall t)).
-destruct (l t); intuition.
-apply is_prefix_pred_correct; intuition.
-Qed.
-
-(** The assumptions about state stack is conserved when we do a goto **)
-Definition goto_past_state :=
- forall s nt,
- match goto_table s nt with
- | Some (exist _ s2 _) =>
- prefix_pred (past_state_of_non_init_state s2)
- (head_states_of_state s)
- | None => True
- end.
-
-Definition is_goto_past_state (_:unit) :=
- forallb (fun s:state =>
- forallb (fun nt =>
- match goto_table s nt with
- | Some (exist _ s2 _) =>
- is_prefix_pred
- (past_state_of_non_init_state s2) (head_states_of_state s)
- | None => true
- end)
- all_list)
- all_list.
-
-Property is_goto_past_state_correct :
- is_goto_past_state () = true -> goto_past_state.
-Proof.
-unfold is_goto_past_state, goto_past_state.
-intros.
-rewrite forallb_forall in H.
-specialize (H s (all_list_forall s)).
-rewrite forallb_forall in H.
-specialize (H nt (all_list_forall nt)).
-destruct (goto_table s nt); intuition.
-destruct s0.
-apply is_prefix_pred_correct; intuition.
-Qed.
-
-(** What states are possible after having popped these symbols from the
- stack, given the annotation of the current state ? **)
-Inductive state_valid_after_pop (s:state):
- list symbol -> list (state -> bool) -> Prop :=
- | state_valid_after_pop_nil1:
- forall p pl, p s = true -> state_valid_after_pop s [] (p::pl)
- | state_valid_after_pop_nil2:
- forall sl, state_valid_after_pop s sl []
- | state_valid_after_pop_cons:
- forall st sq p pl, state_valid_after_pop s sq pl ->
- state_valid_after_pop s (st::sq) (p::pl).
-
-Fixpoint is_state_valid_after_pop
- (state:state) (to_pop:list symbol) annot :=
- match annot, to_pop with
- | [], _ => true
- | p::_, [] => p state
- | p::pl, s::sl => is_state_valid_after_pop state sl pl
- end.
-
-Property is_state_valid_after_pop_complete state sl pl :
- state_valid_after_pop state sl pl -> is_state_valid_after_pop state sl pl = true.
-Proof.
-intro.
-induction H; intuition.
-destruct sl; intuition.
-Qed.
-
-(** A state is valid for reducing a production when :
- - The assumptions on the state are such that we will find the right hand
- side of the production on the stack.
- - We will be able to do a goto after having popped the right hand side.
-**)
-Definition valid_for_reduce (state:state) prod :=
- prefix (prod_rhs_rev prod) (head_symbs_of_state state) /\
- forall state_new,
- state_valid_after_pop state_new
- (prod_rhs_rev prod) (head_states_of_state state) ->
- goto_table state_new (prod_lhs prod) = None ->
- match state_new with
- | Init i => prod_lhs prod = start_nt i
- | Ninit _ => False
- end.
-
-Definition is_valid_for_reduce (state:state) prod:=
- (is_prefix (prod_rhs_rev prod) (head_symbs_of_state state) &&
- forallb (fun state_new =>
- if is_state_valid_after_pop state_new (prod_rhs_rev prod)
- (head_states_of_state state) then
- match goto_table state_new (prod_lhs prod) with
- | Some _ => true
- | None =>
- match state_new with
- | Init i => compare_eqb (prod_lhs prod) (start_nt i)
- | Ninit _ => false
- end
- end
- else
- true)
- all_list)%bool.
-
-Property is_valid_for_reduce_correct (state:state) prod:
- is_valid_for_reduce state prod = true -> valid_for_reduce state prod.
-Proof.
-unfold is_valid_for_reduce, valid_for_reduce.
-intros.
-rewrite Bool.andb_true_iff in H.
-split.
-apply is_prefix_correct.
-intuition.
-intros.
-rewrite forallb_forall in H.
-destruct H.
-specialize (H2 state_new (all_list_forall state_new)).
-rewrite is_state_valid_after_pop_complete, H1 in H2.
-destruct state_new; intuition.
-rewrite compare_eqb_iff in H2; intuition.
-intuition.
-Qed.
-
-(** All the states that does a reduce are valid for reduction **)
-Definition reduce_ok :=
- forall s,
- match action_table s with
- | Lookahead_act awp =>
- forall t, match awp t with
- | Reduce_act p => valid_for_reduce s p
- | _ => True
- end
- | Default_reduce_act p => valid_for_reduce s p
- end.
-
-Definition is_reduce_ok (_:unit) :=
- forallb (fun s =>
- match action_table s with
- | Lookahead_act awp =>
- forallb (fun t =>
- match awp t with
- | Reduce_act p => is_valid_for_reduce s p
- | _ => true
- end)
- all_list
- | Default_reduce_act p => is_valid_for_reduce s p
- end)
- all_list.
-
-Property is_reduce_ok_correct :
- is_reduce_ok () = true -> reduce_ok.
-Proof.
-unfold is_reduce_ok, reduce_ok.
-intros.
-rewrite forallb_forall in H.
-specialize (H s (all_list_forall s)).
-destruct (action_table s).
-apply is_valid_for_reduce_correct; intuition.
-intro.
-rewrite forallb_forall in H.
-specialize (H t (all_list_forall t)).
-destruct (l t); intuition.
-apply is_valid_for_reduce_correct; intuition.
-Qed.
-
-(** The automaton is safe **)
-Definition safe :=
- shift_head_symbs /\ goto_head_symbs /\ shift_past_state /\
- goto_past_state /\ reduce_ok.
-
-Definition is_safe (_:unit) :=
- (is_shift_head_symbs () && is_goto_head_symbs () && is_shift_past_state () &&
- is_goto_past_state () && is_reduce_ok ())%bool.
-
-Property is_safe_correct:
- is_safe () = true -> safe.
-Proof.
-unfold safe, is_safe.
-repeat rewrite Bool.andb_true_iff.
-intuition.
-apply is_shift_head_symbs_correct; assumption.
-apply is_goto_head_symbs_correct; assumption.
-apply is_shift_past_state_correct; assumption.
-apply is_goto_past_state_correct; assumption.
-apply is_reduce_ok_correct; assumption.
-Qed.
-
-End Make.
diff --git a/cparser/PackedStructs.ml b/cparser/PackedStructs.ml
index 3c27f3a9..4c70c7ae 100644
--- a/cparser/PackedStructs.ml
+++ b/cparser/PackedStructs.ml
@@ -418,4 +418,4 @@ let program p =
| _ -> false
end;
Hashtbl.clear byteswapped_fields;
- transf_globdecls (Builtins.environment()) [] p
+ transf_globdecls (Env.initial()) [] p
diff --git a/cparser/Parse.ml b/cparser/Parse.ml
index 154e3dcf..29245083 100644
--- a/cparser/Parse.ml
+++ b/cparser/Parse.ml
@@ -56,22 +56,21 @@ let preprocessed_file transfs name sourcefile =
let text = read_file sourcefile in
let p =
let t = parse_transformations transfs in
- let rec inf = Datatypes.S inf in
+ let log_fuel = Camlcoq.Nat.of_int 50 in
let ast : Cabs.definition list =
- Obj.magic
(match Timing.time "Parsing"
(* The call to Lexer.tokens_stream results in the pre
parsing of the entire file. This is non-negligeabe,
so we cannot use Timing.time2 *)
(fun () ->
- Parser.translation_unit_file inf (Lexer.tokens_stream name text)) ()
+ Parser.translation_unit_file log_fuel (Lexer.tokens_stream name text)) ()
with
- | Parser.Parser.Inter.Fail_pr ->
+ | Parser.MenhirLibParser.Inter.Fail_pr ->
(* Theoretically impossible : implies inconsistencies
between grammars. *)
- Diagnostics.fatal_error Diagnostics.no_loc "internal error while parsing"
- | Parser.Parser.Inter.Timeout_pr -> assert false
- | Parser.Parser.Inter.Parsed_pr (ast, _ ) -> ast) in
+ Diagnostics.fatal_error Diagnostics.no_loc "internal error while parsing"
+ | Parser.MenhirLibParser.Inter.Timeout_pr -> assert false
+ | Parser.MenhirLibParser.Inter.Parsed_pr (ast, _ ) -> ast) in
let p1 = Timing.time "Elaboration" Elab.elab_file ast in
Diagnostics.check_errors ();
Timing.time2 "Emulations" transform_program t p1 name in
diff --git a/cparser/Parser.vy b/cparser/Parser.vy
index 79e3793d..03bfa590 100644
--- a/cparser/Parser.vy
+++ b/cparser/Parser.vy
@@ -15,96 +15,99 @@
%{
-Require Import Cabs.
Require Import List.
+Require Cabs.
%}
-%token<string * cabsloc> VAR_NAME TYPEDEF_NAME OTHER_NAME
-%token<string * cabsloc> PRAGMA
-%token<bool * list char_code * cabsloc> STRING_LITERAL
-%token<constant * cabsloc> CONSTANT
-%token<cabsloc> SIZEOF PTR INC DEC LEFT RIGHT LEQ GEQ EQEQ EQ NEQ LT GT
+%token<Cabs.string * Cabs.loc> VAR_NAME TYPEDEF_NAME OTHER_NAME
+%token<Cabs.string * Cabs.loc> PRAGMA
+%token<bool * list Cabs.char_code * Cabs.loc> STRING_LITERAL
+%token<Cabs.constant * Cabs.loc> CONSTANT
+%token<Cabs.loc> SIZEOF PTR INC DEC LEFT RIGHT LEQ GEQ EQEQ EQ NEQ LT GT
ANDAND BARBAR PLUS MINUS STAR TILDE BANG SLASH PERCENT HAT BAR QUESTION
COLON AND ALIGNOF
-%token<cabsloc> MUL_ASSIGN DIV_ASSIGN MOD_ASSIGN ADD_ASSIGN SUB_ASSIGN
+%token<Cabs.loc> MUL_ASSIGN DIV_ASSIGN MOD_ASSIGN ADD_ASSIGN SUB_ASSIGN
LEFT_ASSIGN RIGHT_ASSIGN AND_ASSIGN XOR_ASSIGN OR_ASSIGN
-%token<cabsloc> LPAREN RPAREN LBRACK RBRACK LBRACE RBRACE DOT COMMA
- SEMICOLON ELLIPSIS TYPEDEF EXTERN STATIC RESTRICT AUTO REGISTER INLINE NORETURN
- CHAR SHORT INT LONG SIGNED UNSIGNED FLOAT DOUBLE CONST VOLATILE VOID
+%token<Cabs.loc> LPAREN RPAREN LBRACK RBRACK LBRACE RBRACE DOT COMMA
+ SEMICOLON ELLIPSIS TYPEDEF EXTERN STATIC RESTRICT AUTO REGISTER INLINE
+ NORETURN CHAR SHORT INT LONG SIGNED UNSIGNED FLOAT DOUBLE CONST VOLATILE VOID
STRUCT UNION ENUM UNDERSCORE_BOOL PACKED ALIGNAS ATTRIBUTE ASM
-%token<cabsloc> CASE DEFAULT IF ELSE SWITCH WHILE DO FOR GOTO CONTINUE BREAK
+%token<Cabs.loc> CASE DEFAULT IF_ ELSE SWITCH WHILE DO FOR GOTO CONTINUE BREAK
RETURN BUILTIN_VA_ARG BUILTIN_OFFSETOF
%token EOF
-%type<expression * cabsloc> primary_expression postfix_expression
+%type<Cabs.expression * Cabs.loc> primary_expression postfix_expression
unary_expression cast_expression multiplicative_expression additive_expression
shift_expression relational_expression equality_expression AND_expression
exclusive_OR_expression inclusive_OR_expression logical_AND_expression
logical_OR_expression conditional_expression assignment_expression
constant_expression expression
-%type<unary_operator * cabsloc> unary_operator
-%type<binary_operator> assignment_operator
-%type<list expression (* Reverse order *)> argument_expression_list
-%type<definition> declaration
-%type<list spec_elem * cabsloc> declaration_specifiers
-%type<list spec_elem> declaration_specifiers_typespec_opt
-%type<list init_name (* Reverse order *)> init_declarator_list
-%type<init_name> init_declarator
-%type<storage * cabsloc> storage_class_specifier
-%type<typeSpecifier * cabsloc> type_specifier struct_or_union_specifier enum_specifier
-%type<structOrUnion * cabsloc> struct_or_union
-%type<list field_group (* Reverse order *)> struct_declaration_list
-%type<field_group> struct_declaration
-%type<list spec_elem * cabsloc> specifier_qualifier_list
-%type<list (option name * option expression) (* Reverse order *)> struct_declarator_list
-%type<option name * option expression> struct_declarator
-%type<list (string * option expression * cabsloc) (* Reverse order *)> enumerator_list
-%type<string * option expression * cabsloc> enumerator
-%type<string * cabsloc> enumeration_constant
-%type<cvspec * cabsloc> type_qualifier type_qualifier_noattr
-%type<funspec * cabsloc> function_specifier
-%type<name> declarator declarator_noattrend direct_declarator
-%type<(decl_type -> decl_type) * cabsloc> pointer
-%type<list cvspec (* Reverse order *)> type_qualifier_list
-%type<list parameter * bool> parameter_type_list
-%type<list parameter (* Reverse order *)> parameter_list
-%type<parameter> parameter_declaration
-%type<list spec_elem * decl_type> type_name
-%type<decl_type> abstract_declarator direct_abstract_declarator
-%type<init_expression> c_initializer
-%type<list (list initwhat * init_expression) (* Reverse order *)> initializer_list
-%type<list initwhat> designation
-%type<list initwhat (* Reverse order *)> designator_list
-%type<initwhat> designator
-%type<statement> statement_dangerous statement_safe
+%type<Cabs.unary_operator * Cabs.loc> unary_operator
+%type<Cabs.binary_operator> assignment_operator
+%type<list Cabs.expression (* Reverse order *)> argument_expression_list
+%type<Cabs.definition> declaration
+%type<list Cabs.spec_elem * Cabs.loc> declaration_specifiers
+%type<list Cabs.spec_elem> declaration_specifiers_typespec_opt
+%type<list Cabs.init_name (* Reverse order *)> init_declarator_list
+%type<Cabs.init_name> init_declarator
+%type<Cabs.storage * Cabs.loc> storage_class_specifier
+%type<Cabs.typeSpecifier * Cabs.loc> type_specifier struct_or_union_specifier enum_specifier
+%type<Cabs.structOrUnion * Cabs.loc> struct_or_union
+%type<list Cabs.field_group (* Reverse order *)> struct_declaration_list
+%type<Cabs.field_group> struct_declaration
+%type<list Cabs.spec_elem * Cabs.loc> specifier_qualifier_list
+%type<list (option Cabs.name * option Cabs.expression) (* Reverse order *)>
+ struct_declarator_list
+%type<option Cabs.name * option Cabs.expression> struct_declarator
+%type<list (Cabs.string * option Cabs.expression * Cabs.loc) (* Reverse order *)>
+ enumerator_list
+%type<Cabs.string * option Cabs.expression * Cabs.loc> enumerator
+%type<Cabs.string * Cabs.loc> enumeration_constant
+%type<Cabs.cvspec * Cabs.loc> type_qualifier type_qualifier_noattr
+%type<Cabs.funspec * Cabs.loc> function_specifier
+%type<Cabs.name> declarator declarator_noattrend direct_declarator
+%type<(Cabs.decl_type -> Cabs.decl_type) * Cabs.loc> pointer
+%type<list Cabs.cvspec (* Reverse order *)> type_qualifier_list
+%type<list Cabs.parameter * bool> parameter_type_list
+%type<list Cabs.parameter (* Reverse order *)> parameter_list
+%type<Cabs.parameter> parameter_declaration
+%type<list Cabs.spec_elem * Cabs.decl_type> type_name
+%type<Cabs.decl_type> abstract_declarator direct_abstract_declarator
+%type<Cabs.init_expression> c_initializer
+%type<list (list Cabs.initwhat * Cabs.init_expression) (* Reverse order *)>
+ initializer_list
+%type<list Cabs.initwhat> designation
+%type<list Cabs.initwhat (* Reverse order *)> designator_list
+%type<Cabs.initwhat> designator
+%type<Cabs.statement> statement_dangerous statement_safe
labeled_statement(statement_safe) labeled_statement(statement_dangerous)
iteration_statement(statement_safe) iteration_statement(statement_dangerous)
compound_statement
-%type<list statement (* Reverse order *)> block_item_list
-%type<statement> block_item expression_statement selection_statement_dangerous
+%type<list Cabs.statement (* Reverse order *)> block_item_list
+%type<Cabs.statement> block_item expression_statement selection_statement_dangerous
selection_statement_safe jump_statement asm_statement
-%type<list definition (* Reverse order *)> translation_unit
-%type<definition> external_declaration function_definition
-%type<list definition> declaration_list
-%type<attribute * cabsloc> attribute_specifier
-%type<list attribute> attribute_specifier_list
-%type<gcc_attribute> gcc_attribute
-%type<list gcc_attribute> gcc_attribute_list
-%type<gcc_attribute_word> gcc_attribute_word
-%type<list string (* Reverse order *)> identifier_list
-%type<list asm_flag> asm_flags
-%type<option string> asm_op_name
-%type<asm_operand> asm_operand
-%type<list asm_operand> asm_operands asm_operands_ne
-%type<list asm_operand * list asm_operand * list asm_flag> asm_arguments
-%type<list cvspec> asm_attributes
-
-%start<list definition> translation_unit_file
+%type<list Cabs.definition (* Reverse order *)> translation_unit
+%type<Cabs.definition> external_declaration function_definition
+%type<list Cabs.definition> declaration_list
+%type<Cabs.attribute * Cabs.loc> attribute_specifier
+%type<list Cabs.attribute> attribute_specifier_list
+%type<Cabs.gcc_attribute> gcc_attribute
+%type<list Cabs.gcc_attribute> gcc_attribute_list
+%type<Cabs.gcc_attribute_word> gcc_attribute_word
+%type<list Cabs.string (* Reverse order *)> identifier_list
+%type<list Cabs.asm_flag> asm_flags
+%type<option Cabs.string> asm_op_name
+%type<Cabs.asm_operand> asm_operand
+%type<list Cabs.asm_operand> asm_operands asm_operands_ne
+%type<list Cabs.asm_operand * list Cabs.asm_operand * list Cabs.asm_flag> asm_arguments
+%type<list Cabs.cvspec> asm_attributes
+
+%start<list Cabs.definition> translation_unit_file
%%
(* Actual grammar *)
@@ -112,12 +115,12 @@ Require Import List.
(* 6.5.1 *)
primary_expression:
| var = VAR_NAME
- { (VARIABLE (fst var), snd var) }
+ { (Cabs.VARIABLE (fst var), snd var) }
| cst = CONSTANT
- { (CONSTANT (fst cst), snd cst) }
+ { (Cabs.CONSTANT (fst cst), snd cst) }
| str = STRING_LITERAL
{ let '((wide, chars), loc) := str in
- (CONSTANT (CONST_STRING wide chars), loc) }
+ (Cabs.CONSTANT (Cabs.CONST_STRING wide chars), loc) }
| loc = LPAREN expr = expression RPAREN
{ (fst expr, loc)}
@@ -126,29 +129,30 @@ postfix_expression:
| expr = primary_expression
{ expr }
| expr = postfix_expression LBRACK index = expression RBRACK
- { (INDEX (fst expr) (fst index), snd expr) }
+ { (Cabs.INDEX (fst expr) (fst index), snd expr) }
| expr = postfix_expression LPAREN args = argument_expression_list RPAREN
- { (CALL (fst expr) (rev' args), snd expr) }
+ { (Cabs.CALL (fst expr) (rev' args), snd expr) }
| expr = postfix_expression LPAREN RPAREN
- { (CALL (fst expr) [], snd expr) }
+ { (Cabs.CALL (fst expr) [], snd expr) }
| loc = BUILTIN_VA_ARG LPAREN expr = assignment_expression COMMA ty = type_name RPAREN
- { (BUILTIN_VA_ARG (fst expr) ty, loc) }
+ { (Cabs.BUILTIN_VA_ARG (fst expr) ty, loc) }
| expr = postfix_expression DOT mem = OTHER_NAME
- { (MEMBEROF (fst expr) (fst mem), snd expr) }
+ { (Cabs.MEMBEROF (fst expr) (fst mem), snd expr) }
| expr = postfix_expression PTR mem = OTHER_NAME
- { (MEMBEROFPTR (fst expr) (fst mem), snd expr) }
+ { (Cabs.MEMBEROFPTR (fst expr) (fst mem), snd expr) }
| expr = postfix_expression INC
- { (UNARY POSINCR (fst expr), snd expr) }
+ { (Cabs.UNARY Cabs.POSINCR (fst expr), snd expr) }
| expr = postfix_expression DEC
- { (UNARY POSDECR (fst expr), snd expr) }
+ { (Cabs.UNARY Cabs.POSDECR (fst expr), snd expr) }
| loc = LPAREN typ = type_name RPAREN LBRACE init = initializer_list RBRACE
- { (CAST typ (COMPOUND_INIT (rev' init)), loc) }
+ { (Cabs.CAST typ (Cabs.COMPOUND_INIT (rev' init)), loc) }
| loc = LPAREN typ = type_name RPAREN LBRACE init = initializer_list COMMA RBRACE
- { (CAST typ (COMPOUND_INIT (rev' init)), loc) }
-| loc = BUILTIN_OFFSETOF LPAREN typ = type_name COMMA id = OTHER_NAME mems = designator_list RPAREN
- { (BUILTIN_OFFSETOF typ ((INFIELD_INIT (fst id))::(rev mems)), loc) }
+ { (Cabs.CAST typ (Cabs.COMPOUND_INIT (rev' init)), loc) }
+| loc = BUILTIN_OFFSETOF LPAREN typ = type_name COMMA id = OTHER_NAME
+ mems = designator_list RPAREN
+ { (Cabs.BUILTIN_OFFSETOF typ ((Cabs.INFIELD_INIT (fst id))::(rev mems)), loc) }
| loc = BUILTIN_OFFSETOF LPAREN typ = type_name COMMA mem = OTHER_NAME RPAREN
- { (BUILTIN_OFFSETOF typ [INFIELD_INIT (fst mem)], loc) }
+ { (Cabs.BUILTIN_OFFSETOF typ [Cabs.INFIELD_INIT (fst mem)], loc) }
(* Semantic value is in reverse order. *)
argument_expression_list:
@@ -162,170 +166,171 @@ unary_expression:
| expr = postfix_expression
{ expr }
| loc = INC expr = unary_expression
- { (UNARY PREINCR (fst expr), loc) }
+ { (Cabs.UNARY Cabs.PREINCR (fst expr), loc) }
| loc = DEC expr = unary_expression
- { (UNARY PREDECR (fst expr), loc) }
+ { (Cabs.UNARY Cabs.PREDECR (fst expr), loc) }
| op = unary_operator expr = cast_expression
- { (UNARY (fst op) (fst expr), snd op) }
+ { (Cabs.UNARY (fst op) (fst expr), snd op) }
| loc = SIZEOF expr = unary_expression
- { (EXPR_SIZEOF (fst expr), loc) }
+ { (Cabs.EXPR_SIZEOF (fst expr), loc) }
| loc = SIZEOF LPAREN typ = type_name RPAREN
- { (TYPE_SIZEOF typ, loc) }
+ { (Cabs.TYPE_SIZEOF typ, loc) }
(* Non-standard *)
| loc = ALIGNOF LPAREN typ = type_name RPAREN
- { (ALIGNOF typ, loc) }
+ { (Cabs.ALIGNOF typ, loc) }
unary_operator:
| loc = AND
- { (ADDROF, loc) }
+ { (Cabs.ADDROF, loc) }
| loc = STAR
- { (MEMOF, loc) }
+ { (Cabs.MEMOF, loc) }
| loc = PLUS
- { (PLUS, loc) }
+ { (Cabs.PLUS, loc) }
| loc = MINUS
- { (MINUS, loc) }
+ { (Cabs.MINUS, loc) }
| loc = TILDE
- { (BNOT, loc) }
+ { (Cabs.BNOT, loc) }
| loc = BANG
- { (NOT, loc) }
+ { (Cabs.NOT, loc) }
(* 6.5.4 *)
cast_expression:
| expr = unary_expression
{ expr }
| loc = LPAREN typ = type_name RPAREN expr = cast_expression
- { (CAST typ (SINGLE_INIT (fst expr)), loc) }
+ { (Cabs.CAST typ (Cabs.SINGLE_INIT (fst expr)), loc) }
(* 6.5.5 *)
multiplicative_expression:
| expr = cast_expression
{ expr }
| expr1 = multiplicative_expression STAR expr2 = cast_expression
- { (BINARY MUL (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.MUL (fst expr1) (fst expr2), snd expr1) }
| expr1 = multiplicative_expression SLASH expr2 = cast_expression
- { (BINARY DIV (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.DIV (fst expr1) (fst expr2), snd expr1) }
| expr1 = multiplicative_expression PERCENT expr2 = cast_expression
- { (BINARY MOD (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.MOD (fst expr1) (fst expr2), snd expr1) }
(* 6.5.6 *)
additive_expression:
| expr = multiplicative_expression
{ expr }
| expr1 = additive_expression PLUS expr2 = multiplicative_expression
- { (BINARY ADD (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.ADD (fst expr1) (fst expr2), snd expr1) }
| expr1 = additive_expression MINUS expr2 = multiplicative_expression
- { (BINARY SUB (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.SUB (fst expr1) (fst expr2), snd expr1) }
(* 6.5.7 *)
shift_expression:
| expr = additive_expression
{ expr }
| expr1 = shift_expression LEFT expr2 = additive_expression
- { (BINARY SHL (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.SHL (fst expr1) (fst expr2), snd expr1) }
| expr1 = shift_expression RIGHT expr2 = additive_expression
- { (BINARY SHR (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.SHR (fst expr1) (fst expr2), snd expr1) }
(* 6.5.8 *)
relational_expression:
| expr = shift_expression
{ expr }
| expr1 = relational_expression LT expr2 = shift_expression
- { (BINARY LT (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.LT (fst expr1) (fst expr2), snd expr1) }
| expr1 = relational_expression GT expr2 = shift_expression
- { (BINARY GT (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.GT (fst expr1) (fst expr2), snd expr1) }
| expr1 = relational_expression LEQ expr2 = shift_expression
- { (BINARY LE (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.LE (fst expr1) (fst expr2), snd expr1) }
| expr1 = relational_expression GEQ expr2 = shift_expression
- { (BINARY GE (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.GE (fst expr1) (fst expr2), snd expr1) }
(* 6.5.9 *)
equality_expression:
| expr = relational_expression
{ expr }
| expr1 = equality_expression EQEQ expr2 = relational_expression
- { (BINARY EQ (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.EQ (fst expr1) (fst expr2), snd expr1) }
| expr1 = equality_expression NEQ expr2 = relational_expression
- { (BINARY NE (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.NE (fst expr1) (fst expr2), snd expr1) }
(* 6.5.10 *)
AND_expression:
| expr = equality_expression
{ expr }
| expr1 = AND_expression AND expr2 = equality_expression
- { (BINARY BAND (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.BAND (fst expr1) (fst expr2), snd expr1) }
(* 6.5.11 *)
exclusive_OR_expression:
| expr = AND_expression
{ expr }
| expr1 = exclusive_OR_expression HAT expr2 = AND_expression
- { (BINARY XOR (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.XOR (fst expr1) (fst expr2), snd expr1) }
(* 6.5.12 *)
inclusive_OR_expression:
| expr = exclusive_OR_expression
{ expr }
| expr1 = inclusive_OR_expression BAR expr2 = exclusive_OR_expression
- { (BINARY BOR (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.BOR (fst expr1) (fst expr2), snd expr1) }
(* 6.5.13 *)
logical_AND_expression:
| expr = inclusive_OR_expression
{ expr }
| expr1 = logical_AND_expression ANDAND expr2 = inclusive_OR_expression
- { (BINARY AND (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.AND (fst expr1) (fst expr2), snd expr1) }
(* 6.5.14 *)
logical_OR_expression:
| expr = logical_AND_expression
{ expr }
| expr1 = logical_OR_expression BARBAR expr2 = logical_AND_expression
- { (BINARY OR (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.OR (fst expr1) (fst expr2), snd expr1) }
(* 6.5.15 *)
conditional_expression:
| expr = logical_OR_expression
{ expr }
-| expr1 = logical_OR_expression QUESTION expr2 = expression COLON expr3 = conditional_expression
- { (QUESTION (fst expr1) (fst expr2) (fst expr3), snd expr1) }
+| expr1 = logical_OR_expression QUESTION expr2 = expression COLON
+ expr3 = conditional_expression
+ { (Cabs.QUESTION (fst expr1) (fst expr2) (fst expr3), snd expr1) }
(* 6.5.16 *)
assignment_expression:
| expr = conditional_expression
{ expr }
| expr1 = unary_expression op = assignment_operator expr2 = assignment_expression
- { (BINARY op (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY op (fst expr1) (fst expr2), snd expr1) }
assignment_operator:
| EQ
- { ASSIGN }
+ { Cabs.ASSIGN }
| MUL_ASSIGN
- { MUL_ASSIGN }
+ { Cabs.MUL_ASSIGN }
| DIV_ASSIGN
- { DIV_ASSIGN }
+ { Cabs.DIV_ASSIGN }
| MOD_ASSIGN
- { MOD_ASSIGN }
+ { Cabs.MOD_ASSIGN }
| ADD_ASSIGN
- { ADD_ASSIGN }
+ { Cabs.ADD_ASSIGN }
| SUB_ASSIGN
- { SUB_ASSIGN }
+ { Cabs.SUB_ASSIGN }
| LEFT_ASSIGN
- { SHL_ASSIGN }
+ { Cabs.SHL_ASSIGN }
| RIGHT_ASSIGN
- { SHR_ASSIGN }
+ { Cabs.SHR_ASSIGN }
| XOR_ASSIGN
- { XOR_ASSIGN }
+ { Cabs.XOR_ASSIGN }
| OR_ASSIGN
- { BOR_ASSIGN }
+ { Cabs.BOR_ASSIGN }
| AND_ASSIGN
- { BAND_ASSIGN }
+ { Cabs.BAND_ASSIGN }
(* 6.5.17 *)
expression:
| expr = assignment_expression
{ expr }
| expr1 = expression COMMA expr2 = assignment_expression
- { (BINARY COMMA (fst expr1) (fst expr2), snd expr1) }
+ { (Cabs.BINARY Cabs.COMMA (fst expr1) (fst expr2), snd expr1) }
(* 6.6 *)
constant_expression:
@@ -335,19 +340,19 @@ constant_expression:
(* 6.7 *)
declaration:
| decspec = declaration_specifiers decls = init_declarator_list SEMICOLON
- { DECDEF (fst decspec, rev' decls) (snd decspec) }
+ { Cabs.DECDEF (fst decspec, rev' decls) (snd decspec) }
| decspec = declaration_specifiers SEMICOLON
- { DECDEF (fst decspec, []) (snd decspec) }
+ { Cabs.DECDEF (fst decspec, []) (snd decspec) }
declaration_specifiers_typespec_opt:
| storage = storage_class_specifier rest = declaration_specifiers_typespec_opt
- { SpecStorage (fst storage)::rest }
+ { Cabs.SpecStorage (fst storage)::rest }
| typ = type_specifier rest = declaration_specifiers_typespec_opt
- { SpecType (fst typ)::rest }
+ { Cabs.SpecType (fst typ)::rest }
| qual = type_qualifier rest = declaration_specifiers_typespec_opt
- { SpecCV (fst qual)::rest }
+ { Cabs.SpecCV (fst qual)::rest }
| func = function_specifier rest = declaration_specifiers_typespec_opt
- { SpecFunction (fst func)::rest }
+ { Cabs.SpecFunction (fst func)::rest }
| /* empty */
{ [] }
@@ -357,16 +362,16 @@ declaration_specifiers_typespec_opt:
specifier. *)
declaration_specifiers:
| storage = storage_class_specifier rest = declaration_specifiers
- { (SpecStorage (fst storage)::fst rest, snd storage) }
+ { (Cabs.SpecStorage (fst storage)::fst rest, snd storage) }
| typ = type_specifier rest = declaration_specifiers_typespec_opt
- { (SpecType (fst typ)::rest, snd typ) }
+ { (Cabs.SpecType (fst typ)::rest, snd typ) }
(* We have to inline type_qualifier in order to avoid a conflict. *)
| qual = type_qualifier_noattr rest = declaration_specifiers
- { (SpecCV (fst qual)::fst rest, snd qual) }
+ { (Cabs.SpecCV (fst qual)::fst rest, snd qual) }
| attr = attribute_specifier rest = declaration_specifiers
- { (SpecCV (CV_ATTR (fst attr))::fst rest, snd attr) }
+ { (Cabs.SpecCV (Cabs.CV_ATTR (fst attr))::fst rest, snd attr) }
| func = function_specifier rest = declaration_specifiers
- { (SpecFunction (fst func)::fst rest, snd func) }
+ { (Cabs.SpecFunction (fst func)::fst rest, snd func) }
init_declarator_list:
| init = init_declarator
@@ -376,71 +381,71 @@ init_declarator_list:
init_declarator:
| name = declarator
- { Init_name name NO_INIT }
+ { Cabs.Init_name name Cabs.NO_INIT }
| name = declarator EQ init = c_initializer
- { Init_name name init }
+ { Cabs.Init_name name init }
(* 6.7.1 *)
storage_class_specifier:
| loc = TYPEDEF
- { (TYPEDEF, loc) }
+ { (Cabs.TYPEDEF, loc) }
| loc = EXTERN
- { (EXTERN, loc) }
+ { (Cabs.EXTERN, loc) }
| loc = STATIC
- { (STATIC, loc) }
+ { (Cabs.STATIC, loc) }
| loc = AUTO
- { (AUTO, loc) }
+ { (Cabs.AUTO, loc) }
| loc = REGISTER
- { (REGISTER, loc) }
+ { (Cabs.REGISTER, loc) }
(* 6.7.2 *)
type_specifier:
| loc = VOID
- { (Tvoid, loc) }
+ { (Cabs.Tvoid, loc) }
| loc = CHAR
- { (Tchar, loc) }
+ { (Cabs.Tchar, loc) }
| loc = SHORT
- { (Tshort, loc) }
+ { (Cabs.Tshort, loc) }
| loc = INT
- { (Tint, loc) }
+ { (Cabs.Tint, loc) }
| loc = LONG
- { (Tlong, loc) }
+ { (Cabs.Tlong, loc) }
| loc = FLOAT
- { (Tfloat, loc) }
+ { (Cabs.Tfloat, loc) }
| loc = DOUBLE
- { (Tdouble, loc) }
+ { (Cabs.Tdouble, loc) }
| loc = SIGNED
- { (Tsigned, loc) }
+ { (Cabs.Tsigned, loc) }
| loc = UNSIGNED
- { (Tunsigned, loc) }
+ { (Cabs.Tunsigned, loc) }
| loc = UNDERSCORE_BOOL
- { (T_Bool, loc) }
+ { (Cabs.T_Bool, loc) }
| spec = struct_or_union_specifier
{ spec }
| spec = enum_specifier
{ spec }
| id = TYPEDEF_NAME
- { (Tnamed (fst id), snd id) }
+ { (Cabs.Tnamed (fst id), snd id) }
(* 6.7.2.1 *)
struct_or_union_specifier:
| str_uni = struct_or_union attrs = attribute_specifier_list id = OTHER_NAME
LBRACE decls = struct_declaration_list RBRACE
- { (Tstruct_union (fst str_uni) (Some (fst id)) (Some (rev' decls)) attrs,
+ { (Cabs.Tstruct_union (fst str_uni) (Some (fst id)) (Some (rev' decls)) attrs,
snd str_uni) }
| str_uni = struct_or_union attrs = attribute_specifier_list
LBRACE decls = struct_declaration_list RBRACE
- { (Tstruct_union (fst str_uni) None (Some (rev' decls)) attrs,
+ { (Cabs.Tstruct_union (fst str_uni) None (Some (rev' decls)) attrs,
snd str_uni) }
| str_uni = struct_or_union attrs = attribute_specifier_list id = OTHER_NAME
- { (Tstruct_union (fst str_uni) (Some (fst id)) None attrs,
+ { (Cabs.Tstruct_union (fst str_uni) (Some (fst id)) None attrs,
snd str_uni) }
struct_or_union:
| loc = STRUCT
- { (STRUCT, loc) }
+ { (Cabs.STRUCT, loc) }
| loc = UNION
- { (UNION, loc) }
+ { (Cabs.UNION, loc) }
struct_declaration_list:
| (* empty *)
@@ -450,20 +455,20 @@ struct_declaration_list:
struct_declaration:
| decspec = specifier_qualifier_list decls = struct_declarator_list SEMICOLON
- { Field_group (fst decspec) (rev' decls) (snd decspec) }
+ { Cabs.Field_group (fst decspec) (rev' decls) (snd decspec) }
(* Extension to C99 grammar needed to parse some GNU header files. *)
| decspec = specifier_qualifier_list SEMICOLON
- { Field_group (fst decspec) [(None,None)] (snd decspec) }
+ { Cabs.Field_group (fst decspec) [(None,None)] (snd decspec) }
specifier_qualifier_list:
| typ = type_specifier rest = specifier_qualifier_list
- { (SpecType (fst typ)::fst rest, snd typ) }
+ { (Cabs.SpecType (fst typ)::fst rest, snd typ) }
| typ = type_specifier
- { ([SpecType (fst typ)], snd typ) }
+ { ([Cabs.SpecType (fst typ)], snd typ) }
| qual = type_qualifier rest = specifier_qualifier_list
- { (SpecCV (fst qual)::fst rest, snd qual) }
+ { (Cabs.SpecCV (fst qual)::fst rest, snd qual) }
| qual = type_qualifier
- { ([SpecCV (fst qual)], snd qual) }
+ { ([Cabs.SpecCV (fst qual)], snd qual) }
struct_declarator_list:
| decl = struct_declarator
@@ -483,18 +488,18 @@ struct_declarator:
enum_specifier:
| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME
LBRACE enum_list = enumerator_list RBRACE
- { (Tenum (Some (fst name)) (Some (rev' enum_list)) attrs, loc) }
+ { (Cabs.Tenum (Some (fst name)) (Some (rev' enum_list)) attrs, loc) }
| loc = ENUM attrs = attribute_specifier_list
LBRACE enum_list = enumerator_list RBRACE
- { (Tenum None (Some (rev' enum_list)) attrs, loc) }
+ { (Cabs.Tenum None (Some (rev' enum_list)) attrs, loc) }
| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME
LBRACE enum_list = enumerator_list COMMA RBRACE
- { (Tenum (Some (fst name)) (Some (rev' enum_list)) attrs, loc) }
+ { (Cabs.Tenum (Some (fst name)) (Some (rev' enum_list)) attrs, loc) }
| loc = ENUM attrs = attribute_specifier_list
LBRACE enum_list = enumerator_list COMMA RBRACE
- { (Tenum None (Some (rev' enum_list)) attrs, loc) }
+ { (Cabs.Tenum None (Some (rev' enum_list)) attrs, loc) }
| loc = ENUM attrs = attribute_specifier_list name = OTHER_NAME
- { (Tenum (Some (fst name)) None attrs, loc) }
+ { (Cabs.Tenum (Some (fst name)) None attrs, loc) }
enumerator_list:
| enum = enumerator
@@ -515,18 +520,18 @@ enumeration_constant:
(* 6.7.3 *)
type_qualifier_noattr:
| loc = CONST
- { (CV_CONST, loc) }
+ { (Cabs.CV_CONST, loc) }
| loc = RESTRICT
- { (CV_RESTRICT, loc) }
+ { (Cabs.CV_RESTRICT, loc) }
| loc = VOLATILE
- { (CV_VOLATILE, loc) }
+ { (Cabs.CV_VOLATILE, loc) }
type_qualifier:
| qual = type_qualifier_noattr
{ qual }
(* Non-standard *)
| attr = attribute_specifier
- { (CV_ATTR (fst attr), snd attr) }
+ { (Cabs.CV_ATTR (fst attr), snd attr) }
(* Non-standard *)
@@ -538,13 +543,13 @@ attribute_specifier_list:
attribute_specifier:
| loc = ATTRIBUTE LPAREN LPAREN attr = gcc_attribute_list RPAREN RPAREN
- { (GCC_ATTR (rev' attr) loc, loc) }
+ { (Cabs.GCC_ATTR (rev' attr) loc, loc) }
| loc = PACKED LPAREN args = argument_expression_list RPAREN
- { (PACKED_ATTR (rev' args) loc, loc) }
+ { (Cabs.PACKED_ATTR (rev' args) loc, loc) }
| loc = ALIGNAS LPAREN args = argument_expression_list RPAREN
- { (ALIGNAS_ATTR (rev' args) loc, loc) }
+ { (Cabs.ALIGNAS_ATTR (rev' args) loc, loc) }
| loc = ALIGNAS LPAREN typ = type_name RPAREN
- { (ALIGNAS_ATTR [ALIGNOF typ] loc, loc) }
+ { (Cabs.ALIGNAS_ATTR [Cabs.ALIGNOF typ] loc, loc) }
gcc_attribute_list:
| a = gcc_attribute
@@ -554,80 +559,81 @@ gcc_attribute_list:
gcc_attribute:
| /* empty */
- { GCC_ATTR_EMPTY }
+ { Cabs.GCC_ATTR_EMPTY }
| w = gcc_attribute_word
- { GCC_ATTR_NOARGS w }
+ { Cabs.GCC_ATTR_NOARGS w }
| w = gcc_attribute_word LPAREN RPAREN
- { GCC_ATTR_ARGS w [] }
+ { Cabs.GCC_ATTR_ARGS w [] }
| w = gcc_attribute_word LPAREN args = argument_expression_list RPAREN
- { GCC_ATTR_ARGS w (rev' args) }
+ { Cabs.GCC_ATTR_ARGS w (rev' args) }
gcc_attribute_word:
| i = OTHER_NAME
- { GCC_ATTR_IDENT (fst i) }
+ { Cabs.GCC_ATTR_IDENT (fst i) }
| CONST
- { GCC_ATTR_CONST }
+ { Cabs.GCC_ATTR_CONST }
| PACKED
- { GCC_ATTR_PACKED }
+ { Cabs.GCC_ATTR_PACKED }
(* 6.7.4 *)
function_specifier:
| loc = INLINE
- { (INLINE, loc) }
+ { (Cabs.INLINE, loc) }
| loc = NORETURN
- { (NORETURN, loc)}
+ { (Cabs.NORETURN, loc)}
(* 6.7.5 *)
declarator:
| decl = declarator_noattrend attrs = attribute_specifier_list
- { match decl with Name name typ attr loc =>
- Name name typ (List.app attr attrs) loc end }
+ { let 'Cabs.Name name typ attr loc := decl in
+ Cabs.Name name typ (List.app attr attrs) loc }
declarator_noattrend:
| decl = direct_declarator
{ decl }
| pt = pointer decl = direct_declarator
- { match decl with Name name typ attr _ =>
- Name name ((fst pt) typ) attr (snd pt) end }
+ { let 'Cabs.Name name typ attr _ := decl in
+ Cabs.Name name ((fst pt) typ) attr (snd pt) }
direct_declarator:
| id = VAR_NAME
- { Name (fst id) JUSTBASE [] (snd id) }
+ { Cabs.Name (fst id) Cabs.JUSTBASE [] (snd id) }
| LPAREN decl = declarator RPAREN
{ decl }
-| decl = direct_declarator LBRACK quallst = type_qualifier_list expr = assignment_expression RBRACK
- { match decl with Name name typ attr loc =>
- Name name (ARRAY typ (rev' quallst) (Some (fst expr))) attr loc end }
+| decl = direct_declarator LBRACK quallst = type_qualifier_list
+ expr = assignment_expression RBRACK
+ { let 'Cabs.Name name typ attr loc := decl in
+ Cabs.Name name (Cabs.ARRAY typ (rev' quallst) (Some (fst expr))) attr loc }
| decl = direct_declarator LBRACK expr = assignment_expression RBRACK
- { match decl with Name name typ attr loc =>
- Name name (ARRAY typ [] (Some (fst expr))) attr loc end }
+ { let 'Cabs.Name name typ attr loc := decl in
+ Cabs.Name name (Cabs.ARRAY typ [] (Some (fst expr))) attr loc }
| decl = direct_declarator LBRACK quallst = type_qualifier_list RBRACK
- { match decl with Name name typ attr loc =>
- Name name (ARRAY typ (rev' quallst) None) attr loc end }
+ { let 'Cabs.Name name typ attr loc := decl in
+ Cabs.Name name (Cabs.ARRAY typ (rev' quallst) None) attr loc }
| decl = direct_declarator LBRACK RBRACK
- { match decl with Name name typ attr loc =>
- Name name (ARRAY typ [] None) attr loc end }
+ { let 'Cabs.Name name typ attr loc := decl in
+ Cabs.Name name (Cabs.ARRAY typ [] None) attr loc }
(*| direct_declarator LBRACK ... STATIC ... RBRACK
| direct_declarator LBRACK STAR RBRACK*)
| decl = direct_declarator LPAREN params = parameter_type_list RPAREN
- { match decl with Name name typ attr loc =>
- Name name (PROTO typ params) attr loc end }
+ { let 'Cabs.Name name typ attr loc := decl in
+ Cabs.Name name (Cabs.PROTO typ params) attr loc }
| decl = direct_declarator LPAREN RPAREN
- { match decl with Name name typ attr loc =>
- Name name (PROTO_OLD typ []) attr loc end }
+ { let 'Cabs.Name name typ attr loc := decl in
+ Cabs.Name name (Cabs.PROTO_OLD typ []) attr loc }
| decl = direct_declarator LPAREN params = identifier_list RPAREN
- { match decl with Name name typ attr loc =>
- Name name (PROTO_OLD typ (rev' params)) attr loc end }
+ { let 'Cabs.Name name typ attr loc := decl in
+ Cabs.Name name (Cabs.PROTO_OLD typ (rev' params)) attr loc }
pointer:
| loc = STAR
- { (fun typ => PTR [] typ, loc) }
+ { (fun typ => Cabs.PTR [] typ, loc) }
| loc = STAR quallst = type_qualifier_list
- { (fun typ => PTR (rev' quallst) typ, loc) }
+ { (fun typ => Cabs.PTR (rev' quallst) typ, loc) }
| loc = STAR pt = pointer
- { (fun typ => PTR [] ((fst pt) typ), loc) }
+ { (fun typ => Cabs.PTR [] ((fst pt) typ), loc) }
| loc = STAR quallst = type_qualifier_list pt = pointer
- { (fun typ => PTR (rev' quallst) ((fst pt) typ), loc) }
+ { (fun typ => Cabs.PTR (rev' quallst) ((fst pt) typ), loc) }
type_qualifier_list:
| qual = type_qualifier
@@ -649,12 +655,12 @@ parameter_list:
parameter_declaration:
| specs = declaration_specifiers decl = declarator
- { match decl with Name name typ attr _ =>
- PARAM (fst specs) (Some name) typ attr (snd specs) end }
+ { match decl with Cabs.Name name typ attr _ =>
+ Cabs.PARAM (fst specs) (Some name) typ attr (snd specs) end }
| specs = declaration_specifiers decl = abstract_declarator
- { PARAM (fst specs) None decl [] (snd specs) }
+ { Cabs.PARAM (fst specs) None decl [] (snd specs) }
| specs = declaration_specifiers
- { PARAM (fst specs) None JUSTBASE [] (snd specs) }
+ { Cabs.PARAM (fst specs) None Cabs.JUSTBASE [] (snd specs) }
identifier_list:
| id = VAR_NAME
@@ -665,13 +671,13 @@ identifier_list:
(* 6.7.6 *)
type_name:
| specqual = specifier_qualifier_list
- { (fst specqual, JUSTBASE) }
+ { (fst specqual, Cabs.JUSTBASE) }
| specqual = specifier_qualifier_list typ = abstract_declarator
{ (fst specqual, typ) }
abstract_declarator:
| pt = pointer
- { (fst pt) JUSTBASE }
+ { (fst pt) Cabs.JUSTBASE }
| pt = pointer typ = direct_abstract_declarator
{ (fst pt) typ }
| typ = direct_abstract_declarator
@@ -680,41 +686,42 @@ abstract_declarator:
direct_abstract_declarator:
| LPAREN typ = abstract_declarator RPAREN
{ typ }
-| typ = direct_abstract_declarator LBRACK cvspec = type_qualifier_list expr = assignment_expression RBRACK
- { ARRAY typ cvspec (Some (fst expr)) }
+| typ = direct_abstract_declarator LBRACK cvspec = type_qualifier_list
+ expr = assignment_expression RBRACK
+ { Cabs.ARRAY typ cvspec (Some (fst expr)) }
| LBRACK cvspec = type_qualifier_list expr = assignment_expression RBRACK
- { ARRAY JUSTBASE cvspec (Some (fst expr)) }
+ { Cabs.ARRAY Cabs.JUSTBASE cvspec (Some (fst expr)) }
| typ = direct_abstract_declarator LBRACK expr = assignment_expression RBRACK
- { ARRAY typ [] (Some (fst expr)) }
+ { Cabs.ARRAY typ [] (Some (fst expr)) }
| LBRACK expr = assignment_expression RBRACK
- { ARRAY JUSTBASE [] (Some (fst expr)) }
+ { Cabs.ARRAY Cabs.JUSTBASE [] (Some (fst expr)) }
| typ = direct_abstract_declarator LBRACK cvspec = type_qualifier_list RBRACK
- { ARRAY typ cvspec None }
+ { Cabs.ARRAY typ cvspec None }
| LBRACK cvspec = type_qualifier_list RBRACK
- { ARRAY JUSTBASE cvspec None }
+ { Cabs.ARRAY Cabs.JUSTBASE cvspec None }
| typ = direct_abstract_declarator LBRACK RBRACK
- { ARRAY typ [] None }
+ { Cabs.ARRAY typ [] None }
| LBRACK RBRACK
- { ARRAY JUSTBASE [] None }
+ { Cabs.ARRAY Cabs.JUSTBASE [] None }
(*| direct_abstract_declarator? LBRACK STAR RBRACK*)
(*| direct_abstract_declarator? LBRACK ... STATIC ... RBRACK*)
| typ = direct_abstract_declarator LPAREN params = parameter_type_list RPAREN
- { PROTO typ params }
+ { Cabs.PROTO typ params }
| LPAREN params = parameter_type_list RPAREN
- { PROTO JUSTBASE params }
+ { Cabs.PROTO Cabs.JUSTBASE params }
| typ = direct_abstract_declarator LPAREN RPAREN
- { PROTO typ ([], false) }
+ { Cabs.PROTO typ ([], false) }
| LPAREN RPAREN
- { PROTO JUSTBASE ([], false) }
+ { Cabs.PROTO Cabs.JUSTBASE ([], false) }
(* 6.7.8 *)
c_initializer:
| expr = assignment_expression
- { SINGLE_INIT (fst expr) }
+ { Cabs.SINGLE_INIT (fst expr) }
| LBRACE init = initializer_list RBRACE
- { COMPOUND_INIT (rev' init) }
+ { Cabs.COMPOUND_INIT (rev' init) }
| LBRACE init = initializer_list COMMA RBRACE
- { COMPOUND_INIT (rev' init) }
+ { Cabs.COMPOUND_INIT (rev' init) }
initializer_list:
| design = designation init = c_initializer
@@ -738,9 +745,9 @@ designator_list:
designator:
| LBRACK expr = constant_expression RBRACK
- { ATINDEX_INIT (fst expr) }
+ { Cabs.ATINDEX_INIT (fst expr) }
| DOT id = OTHER_NAME
- { INFIELD_INIT (fst id) }
+ { Cabs.INFIELD_INIT (fst id) }
(* 6.8 *)
statement_dangerous:
@@ -768,18 +775,18 @@ statement_safe:
(* 6.8.1 *)
labeled_statement(last_statement):
| lbl = OTHER_NAME COLON stmt = last_statement
- { LABEL (fst lbl) stmt (snd lbl) }
+ { Cabs.LABEL (fst lbl) stmt (snd lbl) }
| loc = CASE expr = constant_expression COLON stmt = last_statement
- { CASE (fst expr) stmt loc }
+ { Cabs.CASE (fst expr) stmt loc }
| loc = DEFAULT COLON stmt = last_statement
- { DEFAULT stmt loc }
+ { Cabs.DEFAULT stmt loc }
(* 6.8.2 *)
compound_statement:
| loc = LBRACE lst = block_item_list RBRACE
- { BLOCK (rev' lst) loc }
+ { Cabs.BLOCK (rev' lst) loc }
| loc = LBRACE RBRACE
- { BLOCK [] loc }
+ { Cabs.BLOCK [] loc }
block_item_list:
| stmt = block_item
@@ -789,93 +796,103 @@ block_item_list:
block_item:
| decl = declaration
- { DEFINITION decl }
+ { Cabs.DEFINITION decl }
| stmt = statement_dangerous
{ stmt }
(* Non-standard *)
| p = PRAGMA
- { DEFINITION (PRAGMA (fst p) (snd p)) }
+ { Cabs.DEFINITION (Cabs.PRAGMA (fst p) (snd p)) }
(* 6.8.3 *)
expression_statement:
| expr = expression SEMICOLON
- { COMPUTATION (fst expr) (snd expr) }
+ { Cabs.COMPUTATION (fst expr) (snd expr) }
| loc = SEMICOLON
- { NOP loc }
+ { Cabs.NOP loc }
(* 6.8.4 *)
selection_statement_dangerous:
-| loc = IF LPAREN expr = expression RPAREN stmt = statement_dangerous
- { If (fst expr) stmt None loc }
-| loc = IF LPAREN expr = expression RPAREN stmt1 = statement_safe ELSE stmt2 = statement_dangerous
- { If (fst expr) stmt1 (Some stmt2) loc }
+| loc = IF_ LPAREN expr = expression RPAREN stmt = statement_dangerous
+ { Cabs.If (fst expr) stmt None loc }
+| loc = IF_ LPAREN expr = expression RPAREN stmt1 = statement_safe ELSE
+ stmt2 = statement_dangerous
+ { Cabs.If (fst expr) stmt1 (Some stmt2) loc }
| loc = SWITCH LPAREN expr = expression RPAREN stmt = statement_dangerous
- { SWITCH (fst expr) stmt loc }
+ { Cabs.SWITCH (fst expr) stmt loc }
selection_statement_safe:
-| loc = IF LPAREN expr = expression RPAREN stmt1 = statement_safe ELSE stmt2 = statement_safe
- { If (fst expr) stmt1 (Some stmt2) loc }
+| loc = IF_ LPAREN expr = expression RPAREN stmt1 = statement_safe ELSE
+ stmt2 = statement_safe
+ { Cabs.If (fst expr) stmt1 (Some stmt2) loc }
| loc = SWITCH LPAREN expr = expression RPAREN stmt = statement_safe
- { SWITCH (fst expr) stmt loc }
+ { Cabs.SWITCH (fst expr) stmt loc }
(* 6.8.5 *)
iteration_statement(last_statement):
| loc = WHILE LPAREN expr = expression RPAREN stmt = last_statement
- { WHILE (fst expr) stmt loc }
+ { Cabs.WHILE (fst expr) stmt loc }
| loc = DO stmt = statement_dangerous WHILE LPAREN expr = expression RPAREN SEMICOLON
- { DOWHILE (fst expr) stmt loc }
-| loc = FOR LPAREN expr1 = expression SEMICOLON expr2 = expression SEMICOLON expr3 = expression RPAREN stmt = last_statement
- { FOR (Some (FC_EXP (fst expr1))) (Some (fst expr2)) (Some (fst expr3)) stmt loc }
-| loc = FOR LPAREN decl1 = declaration expr2 = expression SEMICOLON expr3 = expression RPAREN stmt = last_statement
- { FOR (Some (FC_DECL decl1)) (Some (fst expr2)) (Some (fst expr3)) stmt loc }
-| loc = FOR LPAREN SEMICOLON expr2 = expression SEMICOLON expr3 = expression RPAREN stmt = last_statement
- { FOR None (Some (fst expr2)) (Some (fst expr3)) stmt loc }
-| loc = FOR LPAREN expr1 = expression SEMICOLON SEMICOLON expr3 = expression RPAREN stmt = last_statement
- { FOR (Some (FC_EXP (fst expr1))) None (Some (fst expr3)) stmt loc }
-| loc = FOR LPAREN decl1 = declaration SEMICOLON expr3 = expression RPAREN stmt = last_statement
- { FOR (Some (FC_DECL decl1)) None (Some (fst expr3)) stmt loc }
+ { Cabs.DOWHILE (fst expr) stmt loc }
+| loc = FOR LPAREN expr1 = expression SEMICOLON expr2 = expression SEMICOLON
+ expr3 = expression RPAREN stmt = last_statement
+ { Cabs.FOR (Some (Cabs.FC_EXP (fst expr1))) (Some (fst expr2)) (Some (fst expr3)) stmt loc }
+| loc = FOR LPAREN decl1 = declaration expr2 = expression SEMICOLON
+ expr3 = expression RPAREN stmt = last_statement
+ { Cabs.FOR (Some (Cabs.FC_DECL decl1)) (Some (fst expr2)) (Some (fst expr3)) stmt loc }
+| loc = FOR LPAREN SEMICOLON expr2 = expression SEMICOLON expr3 = expression RPAREN
+ stmt = last_statement
+ { Cabs.FOR None (Some (fst expr2)) (Some (fst expr3)) stmt loc }
+| loc = FOR LPAREN expr1 = expression SEMICOLON SEMICOLON expr3 = expression RPAREN
+ stmt = last_statement
+ { Cabs.FOR (Some (Cabs.FC_EXP (fst expr1))) None (Some (fst expr3)) stmt loc }
+| loc = FOR LPAREN decl1 = declaration SEMICOLON expr3 = expression RPAREN
+ stmt = last_statement
+ { Cabs.FOR (Some (Cabs.FC_DECL decl1)) None (Some (fst expr3)) stmt loc }
| loc = FOR LPAREN SEMICOLON SEMICOLON expr3 = expression RPAREN stmt = last_statement
- { FOR None None (Some (fst expr3)) stmt loc }
-| loc = FOR LPAREN expr1 = expression SEMICOLON expr2 = expression SEMICOLON RPAREN stmt = last_statement
- { FOR (Some (FC_EXP (fst expr1))) (Some (fst expr2)) None stmt loc }
-| loc = FOR LPAREN decl1 = declaration expr2 = expression SEMICOLON RPAREN stmt = last_statement
- { FOR (Some (FC_DECL decl1)) (Some (fst expr2)) None stmt loc }
+ { Cabs.FOR None None (Some (fst expr3)) stmt loc }
+| loc = FOR LPAREN expr1 = expression SEMICOLON expr2 = expression SEMICOLON RPAREN
+ stmt = last_statement
+ { Cabs.FOR (Some (Cabs.FC_EXP (fst expr1))) (Some (fst expr2)) None stmt loc }
+| loc = FOR LPAREN decl1 = declaration expr2 = expression SEMICOLON RPAREN
+ stmt = last_statement
+ { Cabs.FOR (Some (Cabs.FC_DECL decl1)) (Some (fst expr2)) None stmt loc }
| loc = FOR LPAREN SEMICOLON expr2 = expression SEMICOLON RPAREN stmt = last_statement
- { FOR None (Some (fst expr2)) None stmt loc }
+ { Cabs.FOR None (Some (fst expr2)) None stmt loc }
| loc = FOR LPAREN expr1 = expression SEMICOLON SEMICOLON RPAREN stmt = last_statement
- { FOR (Some (FC_EXP (fst expr1))) None None stmt loc }
+ { Cabs.FOR (Some (Cabs.FC_EXP (fst expr1))) None None stmt loc }
| loc = FOR LPAREN decl1 = declaration SEMICOLON RPAREN stmt = last_statement
- { FOR (Some (FC_DECL decl1)) None None stmt loc }
+ { Cabs.FOR (Some (Cabs.FC_DECL decl1)) None None stmt loc }
| loc = FOR LPAREN SEMICOLON SEMICOLON RPAREN stmt = last_statement
- { FOR None None None stmt loc }
+ { Cabs.FOR None None None stmt loc }
(* 6.8.6 *)
jump_statement:
| loc = GOTO id = OTHER_NAME SEMICOLON
- { GOTO (fst id) loc }
+ { Cabs.GOTO (fst id) loc }
| loc = CONTINUE SEMICOLON
- { CONTINUE loc }
+ { Cabs.CONTINUE loc }
| loc = BREAK SEMICOLON
- { BREAK loc }
+ { Cabs.BREAK loc }
| loc = RETURN expr = expression SEMICOLON
- { RETURN (Some (fst expr)) loc }
+ { Cabs.RETURN (Some (fst expr)) loc }
| loc = RETURN SEMICOLON
- { RETURN None loc }
+ { Cabs.RETURN None loc }
(* Non-standard *)
asm_statement:
-| loc = ASM attr = asm_attributes LPAREN template = STRING_LITERAL args = asm_arguments RPAREN SEMICOLON
+| loc = ASM attr = asm_attributes LPAREN template = STRING_LITERAL args = asm_arguments
+ RPAREN SEMICOLON
{ let '(wide, chars, _) := template in
let '(outputs, inputs, flags) := args in
- ASM attr wide chars outputs inputs flags loc }
+ Cabs.ASM attr wide chars outputs inputs flags loc }
asm_attributes:
| /* empty */
{ [] }
| CONST attr = asm_attributes
- { CV_CONST :: attr }
+ { Cabs.CV_CONST :: attr }
| VOLATILE attr = asm_attributes
- { CV_VOLATILE :: attr }
+ { Cabs.CV_VOLATILE :: attr }
asm_arguments:
| /* empty */
@@ -897,7 +914,7 @@ asm_operands_ne:
asm_operand:
| n = asm_op_name cstr = STRING_LITERAL LPAREN e = expression RPAREN
- { let '(wide, s, loc) := cstr in ASMOPERAND n wide s (fst e) }
+ { let '(wide, s, loc) := cstr in Cabs.ASMOPERAND n wide s (fst e) }
asm_op_name:
| /* empty */ { None }
@@ -934,7 +951,7 @@ external_declaration:
{ def }
(* Non-standard *)
| p = PRAGMA
- { PRAGMA (fst p) (snd p) }
+ { Cabs.PRAGMA (fst p) (snd p) }
(* 6.9.1 *)
@@ -943,11 +960,11 @@ function_definition:
decl = declarator_noattrend
dlist = declaration_list
stmt = compound_statement
- { FUNDEF (fst specs) decl (List.rev' dlist) stmt (snd specs) }
+ { Cabs.FUNDEF (fst specs) decl (List.rev' dlist) stmt (snd specs) }
| specs = declaration_specifiers
decl = declarator
stmt = compound_statement
- { FUNDEF (fst specs) decl [] stmt (snd specs) }
+ { Cabs.FUNDEF (fst specs) decl [] stmt (snd specs) }
declaration_list:
| d = declaration
diff --git a/cparser/Rename.ml b/cparser/Rename.ml
index eb31eaf0..64412194 100644
--- a/cparser/Rename.ml
+++ b/cparser/Rename.ml
@@ -246,7 +246,7 @@ let rec globdecls env accu = function
(* Reserve names of builtins *)
let reserve_builtins () =
- List.fold_left enter_public empty_env (Builtins.identifiers())
+ List.fold_left enter_public empty_env (Env.initial_identifiers())
(* Reserve global declarations with public visibility *)
diff --git a/cparser/Transform.ml b/cparser/Transform.ml
index 0a2ce3bb..349a3155 100644
--- a/cparser/Transform.ml
+++ b/cparser/Transform.ml
@@ -227,4 +227,4 @@ let program
in
transf_globdecls env' ({g with gdesc = desc'} :: accu) gl
- in transf_globdecls (Builtins.environment()) [] p
+ in transf_globdecls (Env.initial()) [] p
diff --git a/cparser/Unblock.ml b/cparser/Unblock.ml
index 66b497cc..d25f70c6 100644
--- a/cparser/Unblock.ml
+++ b/cparser/Unblock.ml
@@ -399,4 +399,4 @@ let rec unblock_glob env accu = function
let program p =
next_scope_id := 0;
{gloc = no_loc; gdesc = Gdecl(Storage_extern, debug_id, debug_ty, None)} ::
- unblock_glob (Builtins.environment()) [] p
+ unblock_glob (Env.initial()) [] p
diff --git a/cparser/pre_parser.mly b/cparser/pre_parser.mly
index 71eaf419..669ecf5e 100644
--- a/cparser/pre_parser.mly
+++ b/cparser/pre_parser.mly
@@ -43,13 +43,13 @@
%}
%token<string> PRE_NAME
-%token<string * Pre_parser_aux.identifier_type ref * Cabs.cabsloc>
+%token<string * Pre_parser_aux.identifier_type ref * Cabs.loc>
VAR_NAME TYPEDEF_NAME
-%token<Cabs.constant * Cabs.cabsloc> CONSTANT
-%token<bool * int64 list * Cabs.cabsloc> STRING_LITERAL
-%token<string * Cabs.cabsloc> PRAGMA
+%token<Cabs.constant * Cabs.loc> CONSTANT
+%token<bool * int64 list * Cabs.loc> STRING_LITERAL
+%token<string * Cabs.loc> PRAGMA
-%token<Cabs.cabsloc> SIZEOF PTR INC DEC LEFT RIGHT LEQ GEQ EQEQ EQ NEQ LT GT
+%token<Cabs.loc> SIZEOF PTR INC DEC LEFT RIGHT LEQ GEQ EQEQ EQ NEQ LT GT
ANDAND BARBAR PLUS MINUS STAR TILDE BANG SLASH PERCENT HAT BAR QUESTION
COLON AND MUL_ASSIGN DIV_ASSIGN MOD_ASSIGN ADD_ASSIGN SUB_ASSIGN LEFT_ASSIGN
RIGHT_ASSIGN AND_ASSIGN XOR_ASSIGN OR_ASSIGN LPAREN RPAREN LBRACK RBRACK
diff --git a/debug/DebugInformation.ml b/debug/DebugInformation.ml
index 3498a779..f9684355 100644
--- a/debug/DebugInformation.ml
+++ b/debug/DebugInformation.ml
@@ -123,7 +123,7 @@ let insert_type ty =
| TNamed (id,_) ->
let typ = try
let _,t =
- List.find (fun a -> fst a = id.name) CBuiltins.builtins.Builtins.typedefs in
+ List.find (fun a -> fst a = id.name) CBuiltins.builtins.builtin_typedefs in
Some (attr_aux t)
with Not_found -> None in
let t = {
diff --git a/debug/DwarfPrinter.ml b/debug/DwarfPrinter.ml
index bbfcf311..9a24041b 100644
--- a/debug/DwarfPrinter.ml
+++ b/debug/DwarfPrinter.ml
@@ -268,7 +268,7 @@ module DwarfPrinter(Target: DWARF_TARGET):
(* Print the debug_abbrev section using the previous computed abbreviations*)
let print_abbrev oc =
let abbrevs = Hashtbl.fold (fun s i acc -> (s,i)::acc) abbrev_mapping [] in
- let abbrevs = List.sort (fun (_,a) (_,b) -> Pervasives.compare a b) abbrevs in
+ let abbrevs = List.sort (fun (_,a) (_,b) -> compare a b) abbrevs in
section oc Section_debug_abbrev;
print_label oc !abbrev_start_addr;
List.iter (fun (s,id) ->
@@ -685,7 +685,7 @@ module DwarfPrinter(Target: DWARF_TARGET):
print_label oc line_start;
list_opt s (fun () ->
section oc Section_debug_str;
- let s = List.sort (fun (a,_) (b,_) -> Pervasives.compare a b) s in
+ let s = List.sort (fun (a,_) (b,_) -> compare a b) s in
List.iter (fun (id,s) ->
print_label oc (loc_to_label id);
fprintf oc " .asciz %S\n" s) s)
diff --git a/doc/ccomp.1 b/doc/ccomp.1
index 7ccf97c8..f4919867 100644
--- a/doc/ccomp.1
+++ b/doc/ccomp.1
@@ -125,7 +125,8 @@ Enabled by default.
.TP
.B \-O0
Turn off most optimizations.
-Synonymous to \fB\-fno\-const\-prop\fP \fB\-fno\-cse\fP \fB\-fno\-redundancy\fP \fB\-fno\-tailcalls\fP.
+Synonymous to \fB\-fno\-const\-prop\fP \fB\-fno\-cse\fP \fB\-fno\-if\-conversion\fP
+\fB\-fno\-inline\fP \fB\-fno\-redundancy\fP \fB\-fno\-tailcalls\fP.
.
.TP
.BR \-O1 ", " \-O2 ", " \-O3
@@ -136,6 +137,13 @@ Synonymous for \fB\-O\fP.
Optimize for code size in preference to code speed.
.
.TP
+.B \-Obranchless
+Optimize to generate fewer conditional branches and use branch-free
+instruction sequences instead. When \fB-fif\-conversion\fP is
+enabled, the conversion is peformed aggressively even if the resulting
+code is less performant.
+.
+.TP
.BR \-fconst\-prop ", " \-fno\-const\-prop
Turn on/off global constant propagation.
Enabled by default.
@@ -146,6 +154,11 @@ Turn on/off common subexpression elimination.
Enabled by default.
.
.TP
+.BR \-fif\-conversion ", " \-fno\-if\-conversion
+Turn on/off generation of conditional moves.
+Enabled by default.
+.
+.TP
.BR \-finline ", " \-fno\-inline
Turn on/off inlining of functions.
Enabled by default.
@@ -433,6 +446,11 @@ Enabled by default.
Declarations which do not declare anything.
Enabled by default.
.sp
+\fInon\-linear\-cond\-expr\fP:
+Conditional expression that may not be optimized to branchless code.
+Only issued in \fB-Obranchless\fP mode.
+Disabled by default.
+.sp
\fIpointer\-type\-mismatch\fP:
Use of incompatible pointer types in conditional expressions.
Enabled by default.
diff --git a/driver/Clflags.ml b/driver/Clflags.ml
index 00bbc4d5..cf1220d1 100644
--- a/driver/Clflags.ml
+++ b/driver/Clflags.ml
@@ -30,7 +30,7 @@ let option_fredundancy = ref true
let option_fpostpass = ref true
let option_fpostpass_sched = ref "list"
let option_fifconversion = ref true
-let option_ffavor_branchless = ref false
+let option_Obranchless = ref false
let option_falignfunctions = ref (None: int option)
let option_falignbranchtargets = ref 0
let option_faligncondbranchs = ref 0
diff --git a/driver/Driver.ml b/driver/Driver.ml
index 9748ebf6..288bb436 100644
--- a/driver/Driver.ml
+++ b/driver/Driver.ml
@@ -188,6 +188,8 @@ Processing options:
-O1 Perform all optimization passes except scheduling
-O2 -O3 Synonymous for -O
-Os Optimize for code size in preference to code speed
+ -Obranchless Optimize to generate fewer conditional branches; try to produce
+ branch-free instruction sequences as much as possible
-ftailcalls Optimize function calls in tail position [on]
-fconst-prop Perform global constant propagation [on]
-ffloat-const-prop <n> Control constant propagation of floats
@@ -201,8 +203,6 @@ Processing options:
-finline-functions-called-once Integrate functions only required by their
single caller [on]
-fif-conversion Perform if-conversion (generation of conditional moves) [on]
- -ffavor-branchless Favor the generation of branch-free instruction sequences,
- even when possibly more costly than the default [off]
Code generation options: (use -fno-<opt> to turn off -f<opt>)
-ffpu Use FP registers for some integer operations [on]
-fsmall-data <n> Set maximal size <n> for allocation in small data area
@@ -312,10 +312,10 @@ let cmdline_actions =
_Regexp "-O1", Self (fun _ -> set_all optimization_options (); option_fpostpass := false);
_Regexp "-O[123]$", Unit (set_all optimization_options);
Exact "-Os", Set option_Osize;
+ Exact "-Obranchless", Set option_Obranchless;
Exact "-fsmall-data", Integer(fun n -> option_small_data := n);
Exact "-fsmall-const", Integer(fun n -> option_small_const := n);
Exact "-ffloat-const-prop", Integer(fun n -> option_ffloatconstprop := n);
- Exact "-ffavor-branchless", Set option_ffavor_branchless;
Exact "-falign-functions", Integer(fun n -> check_align n; option_falignfunctions := Some n);
Exact "-falign-branch-targets", Integer(fun n -> check_align n; option_falignbranchtargets := n);
Exact "-falign-cond-branches", Integer(fun n -> check_align n; option_faligncondbranchs := n);] @
diff --git a/driver/Frontend.ml b/driver/Frontend.ml
index d7162865..2584db90 100644
--- a/driver/Frontend.ml
+++ b/driver/Frontend.ml
@@ -119,7 +119,7 @@ let init () =
| "mppa_k1c" -> Machine.mppa_k1c
| _ -> assert false
end;
- Builtins.set C2C.builtins;
+ Env.set_builtins C2C.builtins;
Cutil.declare_attributes C2C.attributes;
CPragmas.initialize()
diff --git a/driver/Interp.ml b/driver/Interp.ml
index 6760e76c..a6841460 100644
--- a/driver/Interp.ml
+++ b/driver/Interp.ml
@@ -15,7 +15,7 @@
open Format
open Camlcoq
open AST
-open Integers
+open !Integers
open Values
open Memory
open Globalenvs
@@ -145,7 +145,7 @@ let print_state p (prog, ge, s) =
let compare_mem m1 m2 =
(* assumes nextblocks were already compared equal *)
(* should permissions be taken into account? *)
- Pervasives.compare m1.Mem.mem_contents m2.Mem.mem_contents
+ compare m1.Mem.mem_contents m2.Mem.mem_contents
(* Comparing continuations *)
diff --git a/exportclight/Clightgen.ml b/exportclight/Clightgen.ml
index 1eb4fe03..4209975a 100644
--- a/exportclight/Clightgen.ml
+++ b/exportclight/Clightgen.ml
@@ -45,12 +45,7 @@ let compile_c_ast sourcename csyntax ofile =
| Errors.Error msg ->
fatal_error loc "%a" print_error msg in
(* Dump Clight in C syntax if requested *)
- if !option_dclight then begin
- let ofile = Filename.chop_suffix sourcename ".c" ^ ".light.c" in
- let oc = open_out ofile in
- PrintClight.print_program (Format.formatter_of_out_channel oc) clight;
- close_out oc
- end;
+ PrintClight.print_if clight;
(* Print Clight in Coq syntax *)
let oc = open_out ofile in
ExportClight.print_program (Format.formatter_of_out_channel oc)
@@ -60,6 +55,12 @@ let compile_c_ast sourcename csyntax ofile =
(* From C source to Clight *)
let compile_c_file sourcename ifile ofile =
+ let set_dest dst opt ext =
+ dst := if !opt then Some (output_filename sourcename ".c" ext)
+ else None in
+ set_dest Cprint.destination option_dparse ".parsed.c";
+ set_dest PrintCsyntax.destination option_dcmedium ".compcert.c";
+ set_dest PrintClight.destination option_dclight ".light.c";
compile_c_ast sourcename (parse_c_file sourcename ifile) ofile
let output_filename sourcename suff =
@@ -74,7 +75,10 @@ let process_c_file sourcename =
if !option_E then begin
preprocess sourcename "-"
end else begin
- let preproname = Driveraux.tmp_file ".i" in
+ let preproname = if !option_dprepro then
+ Driveraux.output_filename sourcename ".c" ".i"
+ else
+ Driveraux.tmp_file ".i" in
preprocess sourcename preproname;
compile_c_file sourcename preproname ofile
end
@@ -100,9 +104,11 @@ Processing options:
prepro_help ^
language_support_help ^
{|Tracing options:
+ -dprepro Save C file after preprocessing in <file>.i
-dparse Save C file after parsing and elaboration in <file>.parsed.c
-dc Save generated Compcert C in <file>.compcert.c
-dclight Save generated Clight in <file>.light.c
+ -dall Save all generated intermediate files in <file>.<ext>
|} ^
general_help ^
warning_help
@@ -142,9 +148,16 @@ let cmdline_actions =
(* Preprocessing options *)
@ prepro_actions @
(* Tracing options *)
- [ Exact "-dparse", Set option_dparse;
- Exact "-dc", Set option_dcmedium;
- Exact "-dclight", Set option_dclight;]
+ [ Exact "-dprepro", Set option_dprepro;
+ Exact "-dparse", Set option_dparse;
+ Exact "-dc", Set option_dcmedium;
+ Exact "-dclight", Set option_dclight;
+ Exact "-dall", Self (fun _ ->
+ option_dprepro := true;
+ option_dparse := true;
+ option_dcmedium := true;
+ option_dclight := true;);
+ ]
@ general_options
(* Diagnostic options *)
@ warning_options
diff --git a/extraction/extraction.v b/extraction/extraction.v
index 265a5967..e4c1cb25 100644
--- a/extraction/extraction.v
+++ b/extraction/extraction.v
@@ -34,7 +34,6 @@ Require Clight.
Require Compiler.
Require Parser.
Require Initializers.
-Require Int31.
Require Asmaux.
(* Standard lib *)
@@ -141,7 +140,7 @@ Extract Constant Compiler.time => "Timing.time_coq".
(*Extraction Inline Compiler.apply_total Compiler.apply_partial.*)
(* Cabs *)
-Extract Constant Cabs.cabsloc =>
+Extract Constant Cabs.loc =>
"{ lineno : int;
filename: string;
byteno: int;
@@ -150,15 +149,6 @@ Extract Constant Cabs.cabsloc =>
Extract Inlined Constant Cabs.string => "String.t".
Extract Constant Cabs.char_code => "int64".
-(* Int31 *)
-Extract Inductive Int31.digits => "bool" [ "false" "true" ].
-Extract Inductive Int31.int31 => "int" [ "Camlcoq.Int31.constr" ] "Camlcoq.Int31.destr".
-Extract Constant Int31.twice => "Camlcoq.Int31.twice".
-Extract Constant Int31.twice_plus_one => "Camlcoq.Int31.twice_plus_one".
-Extract Constant Int31.compare31 => "Camlcoq.Int31.compare".
-Extract Constant Int31.On => "0".
-Extract Constant Int31.In => "1".
-
(* Processor-specific extraction directives *)
Load extractionMachdep.
@@ -185,9 +175,10 @@ Separate Extraction
Cexec.do_initial_state Cexec.do_step Cexec.at_final_state
Ctypes.merge_attributes Ctypes.remove_attributes Ctypes.build_composite_env
Initializers.transl_init Initializers.constval
- Csyntax.Eindex Csyntax.Epreincr
+ Csyntax.Eindex Csyntax.Epreincr Csyntax.Eselection
Ctyping.typecheck_program
Ctyping.epostincr Ctyping.epostdecr Ctyping.epreincr Ctyping.epredecr
+ Ctyping.eselection
Ctypes.make_program
Clight.type_of_function
Conventions1.callee_save_type Conventions1.is_float_reg
diff --git a/flocq/IEEE754/Binary.v b/flocq/IEEE754/Binary.v
index 0ec3a297..ac38c761 100644
--- a/flocq/IEEE754/Binary.v
+++ b/flocq/IEEE754/Binary.v
@@ -1839,6 +1839,127 @@ now rewrite <- cond_Zopp_negb.
now destruct y as [ | | | ].
Qed.
+(** Fused Multiply-Add *)
+
+Definition Bfma_szero m (x y z: binary_float) : bool :=
+ let s_xy := xorb (Bsign x) (Bsign y) in (* sign of product x*y *)
+ if Bool.eqb s_xy (Bsign z) then s_xy
+ else match m with mode_DN => true | _ => false end.
+
+Definition Bfma fma_nan m (x y z: binary_float) :=
+ match x, y with
+ | B754_nan _ _ _, _ | _, B754_nan _ _ _
+ | B754_infinity _, B754_zero _
+ | B754_zero _, B754_infinity _ =>
+ (* Multiplication produces NaN *)
+ build_nan (fma_nan x y z)
+ | B754_infinity sx, B754_infinity sy
+ | B754_infinity sx, B754_finite sy _ _ _
+ | B754_finite sx _ _ _, B754_infinity sy =>
+ let s := xorb sx sy in
+ (* Multiplication produces infinity with sign [s] *)
+ match z with
+ | B754_nan _ _ _ => build_nan (fma_nan x y z)
+ | B754_infinity sz =>
+ if Bool.eqb s sz then z else build_nan (fma_nan x y z)
+ | _ => B754_infinity s
+ end
+ | B754_finite sx _ _ _, B754_zero sy
+ | B754_zero sx, B754_finite sy _ _ _
+ | B754_zero sx, B754_zero sy =>
+ (* Multiplication produces zero *)
+ match z with
+ | B754_nan _ _ _ => build_nan (fma_nan x y z)
+ | B754_zero _ => B754_zero (Bfma_szero m x y z)
+ | _ => z
+ end
+ | B754_finite sx mx ex _, B754_finite sy my ey _ =>
+ (* Multiplication produces a finite, non-zero result *)
+ match z with
+ | B754_nan _ _ _ => build_nan (fma_nan x y z)
+ | B754_infinity sz => z
+ | B754_zero _ =>
+ let X := Float radix2 (cond_Zopp sx (Zpos mx)) ex in
+ let Y := Float radix2 (cond_Zopp sy (Zpos my)) ey in
+ let '(Float _ mr er) := Fmult X Y in
+ binary_normalize m mr er (Bfma_szero m x y z)
+ | B754_finite sz mz ez _ =>
+ let X := Float radix2 (cond_Zopp sx (Zpos mx)) ex in
+ let Y := Float radix2 (cond_Zopp sy (Zpos my)) ey in
+ let Z := Float radix2 (cond_Zopp sz (Zpos mz)) ez in
+ let '(Float _ mr er) := Fplus (Fmult X Y) Z in
+ binary_normalize m mr er (Bfma_szero m x y z)
+ end
+ end.
+
+Theorem Bfma_correct:
+ forall fma_nan m x y z,
+ let res := (B2R x * B2R y + B2R z)%R in
+ is_finite x = true ->
+ is_finite y = true ->
+ is_finite z = true ->
+ if Rlt_bool (Rabs (round radix2 fexp (round_mode m) res)) (bpow radix2 emax) then
+ B2R (Bfma fma_nan m x y z) = round radix2 fexp (round_mode m) res /\
+ is_finite (Bfma fma_nan m x y z) = true /\
+ Bsign (Bfma fma_nan m x y z) =
+ match Rcompare res 0 with
+ | Eq => Bfma_szero m x y z
+ | Lt => true
+ | Gt => false
+ end
+ else
+ B2FF (Bfma fma_nan m x y z) = binary_overflow m (Rlt_bool res 0).
+Proof.
+ intros. pattern (Bfma fma_nan m x y z).
+ match goal with |- ?p ?x => set (PROP := p) end.
+ set (szero := Bfma_szero m x y z).
+ assert (BINORM: forall mr er, F2R (Float radix2 mr er) = res ->
+ PROP (binary_normalize m mr er szero)).
+ { intros mr er E.
+ specialize (binary_normalize_correct m mr er szero).
+ change (FLT_exp (3 - emax - prec) prec) with fexp. rewrite E. tauto.
+ }
+ set (add_zero :=
+ match z with
+ | B754_nan _ _ _ => build_nan (fma_nan x y z)
+ | B754_zero sz => B754_zero szero
+ | _ => z
+ end).
+ assert (ADDZERO: B2R x = 0%R \/ B2R y = 0%R -> PROP add_zero).
+ {
+ intros Z.
+ assert (RES: res = B2R z).
+ { unfold res. destruct Z as [E|E]; rewrite E, ?Rmult_0_l, ?Rmult_0_r, Rplus_0_l; auto. }
+ unfold PROP, add_zero; destruct z as [ sz | sz | sz plz | sz mz ez Bz]; try discriminate.
+ - simpl in RES; rewrite RES; rewrite round_0 by apply valid_rnd_round_mode.
+ rewrite Rlt_bool_true. split. reflexivity. split. reflexivity.
+ rewrite Rcompare_Eq by auto. reflexivity.
+ rewrite Rabs_R0; apply bpow_gt_0.
+ - rewrite RES, round_generic, Rlt_bool_true.
+ split. reflexivity. split. reflexivity.
+ unfold B2R. destruct sz.
+ rewrite Rcompare_Lt. auto. apply F2R_lt_0. reflexivity.
+ rewrite Rcompare_Gt. auto. apply F2R_gt_0. reflexivity.
+ apply abs_B2R_lt_emax. apply valid_rnd_round_mode. apply generic_format_B2R.
+ }
+ destruct x as [ sx | sx | sx plx | sx mx ex Bx];
+ destruct y as [ sy | sy | sy ply | sy my ey By];
+ try discriminate.
+- apply ADDZERO; auto.
+- apply ADDZERO; auto.
+- apply ADDZERO; auto.
+- destruct z as [ sz | sz | sz plz | sz mz ez Bz]; try discriminate; unfold Bfma.
++ set (X := Float radix2 (cond_Zopp sx (Zpos mx)) ex).
+ set (Y := Float radix2 (cond_Zopp sy (Zpos my)) ey).
+ destruct (Fmult X Y) as [mr er] eqn:FRES.
+ apply BINORM. unfold res. rewrite <- FRES, F2R_mult, Rplus_0_r. auto.
++ set (X := Float radix2 (cond_Zopp sx (Zpos mx)) ex).
+ set (Y := Float radix2 (cond_Zopp sy (Zpos my)) ey).
+ set (Z := Float radix2 (cond_Zopp sz (Zpos mz)) ez).
+ destruct (Fplus (Fmult X Y) Z) as [mr er] eqn:FRES.
+ apply BINORM. unfold res. rewrite <- FRES, F2R_plus, F2R_mult. auto.
+Qed.
+
(** Division *)
Definition Fdiv_core_binary m1 e1 m2 e2 :=
diff --git a/lib/Camlcoq.ml b/lib/Camlcoq.ml
index d94e3582..66322efb 100644
--- a/lib/Camlcoq.ml
+++ b/lib/Camlcoq.ml
@@ -335,54 +335,3 @@ let coqfloat32_of_camlfloat f =
Float32.of_bits(coqint_of_camlint(Int32.bits_of_float f))
let camlfloat_of_coqfloat32 f =
Int32.float_of_bits(camlint_of_coqint(Float32.to_bits f))
-
-(* Int31 *)
-
-module Int31 = struct
-
-(*
-let constr (b30,b29,b28,b27,b26,b25,b24,
- b23,b22,b21,b20,b19,b18,b17,b16,
- b15,b14,b13,b12,b11,b10,b9,b8,
- b7,b6,b5,b4,b3,b2,b1,b0) =
- let f i b accu = if b then accu + (1 lsl i) else accu in
- f 30 b30 (f 29 b29 (f 28 b28 (f 27 b27 (f 26 b26 (f 25 b25 (f 24 b24
- (f 23 b23 (f 22 b22 (f 21 b21 (f 20 b20 (f 19 b19 (f 18 b18 (f 17 b17 (f 16 b16
- (f 15 b15 (f 14 b14 (f 13 b13 (f 12 b12 (f 11 b11 (f 10 b10 (f 9 b9 (f 8 b8
- (f 7 b7 (f 6 b6 (f 5 b5 (f 4 b4 (f 3 b3 (f 2 b2 (f 1 b1 (f 0 b0 0))))))))))))))))))))))))))))))
-*)
-
-let constr (b30,b29,b28,b27,b26,b25,b24,
- b23,b22,b21,b20,b19,b18,b17,b16,
- b15,b14,b13,b12,b11,b10,b9,b8,
- b7,b6,b5,b4,b3,b2,b1,b0) =
- let f i b = if b then 1 lsl i else 0 in
- f 30 b30 + f 29 b29 + f 28 b28 + f 27 b27 + f 26 b26 + f 25 b25 + f 24 b24 +
- f 23 b23 + f 22 b22 + f 21 b21 + f 20 b20 + f 19 b19 + f 18 b18 + f 17 b17 + f 16 b16 +
- f 15 b15 + f 14 b14 + f 13 b13 + f 12 b12 + f 11 b11 + f 10 b10 + f 9 b9 + f 8 b8 +
- f 7 b7 + f 6 b6 + f 5 b5 + f 4 b4 + f 3 b3 + f 2 b2 + f 1 b1 + f 0 b0
-
-let destr f n =
- let b i = n land (1 lsl i) <> 0 in
- f (b 30) (b 29) (b 28) (b 27) (b 26) (b 25) (b 24)
- (b 23) (b 22) (b 21) (b 20) (b 19) (b 18) (b 17) (b 16)
- (b 15) (b 14) (b 13) (b 12) (b 11) (b 10) (b 9) (b 8)
- (b 7) (b 6) (b 5) (b 4) (b 3) (b 2) (b 1) (b 0)
-
-let twice n =
- (n lsl 1) land 0x7FFFFFFF
-
-let twice_plus_one n =
- ((n lsl 1) land 0x7FFFFFFF) lor 1
-
-let compare (x:int) (y:int) =
- if x = y then Datatypes.Eq
- else begin
- let sx = x < 0 and sy = y < 0 in
- if sx = sy then
- (if x < y then Datatypes.Lt else Datatypes.Gt)
- else
- (if sx then Datatypes.Gt else Datatypes.Lt)
- end
-
-end
diff --git a/lib/Floats.v b/lib/Floats.v
index 9540303b..13350dd0 100644
--- a/lib/Floats.v
+++ b/lib/Floats.v
@@ -24,6 +24,7 @@ Require Import Program.
Require Archi.
Close Scope R_scope.
+Open Scope Z_scope.
Definition float := binary64. (**r the type of IEE754 double-precision FP numbers *)
Definition float32 := binary32. (**r the type of IEE754 single-precision FP numbers *)
@@ -93,21 +94,53 @@ Proof.
destruct x as [[]|]; simpl; intros; discriminate.
Qed.
-(** Relation between number of bits and base-2 logarithm *)
+(** Normalization of NaN payloads *)
-Lemma digits2_log2:
- forall p, Z.pos (Digits.digits2_pos p) = Z.succ (Z.log2 (Z.pos p)).
+Lemma normalized_nan: forall prec n p,
+ Z.of_nat n = prec - 1 -> 1 < prec ->
+ nan_pl prec (Z.to_pos (P_mod_two_p p n)) = true.
Proof.
- assert (E: forall p, Digits.digits2_pos p = Pos.size p).
- { induction p; simpl; rewrite ?IHp; auto. }
- intros p. rewrite E.
- destruct p; simpl; rewrite ?Pos.add_1_r; reflexivity.
+ intros. unfold nan_pl. apply Z.ltb_lt. rewrite Digits.Zpos_digits2_pos.
+ set (p' := P_mod_two_p p n).
+ assert (A: 0 <= p' < 2 ^ Z.of_nat n).
+ { rewrite <- two_power_nat_equiv; apply P_mod_two_p_range. }
+ assert (B: Digits.Zdigits radix2 p' <= prec - 1).
+ { apply Digits.Zdigits_le_Zpower. rewrite <- H. rewrite Z.abs_eq; tauto. }
+ destruct (zeq p' 0).
+- rewrite e. simpl; auto.
+- rewrite Z2Pos.id by omega. omega.
Qed.
+(** Transform a Nan payload to a quiet Nan payload. *)
+
+Definition quiet_nan_64_payload (p: positive) :=
+ Z.to_pos (P_mod_two_p (Pos.lor p ((iter_nat xO 51 1%positive))) 52%nat).
+
+Lemma quiet_nan_64_proof: forall p, nan_pl 53 (quiet_nan_64_payload p) = true.
+Proof. intros; apply normalized_nan; auto; omega. Qed.
+
+Definition quiet_nan_64 (sp: bool * positive) : {x :float | is_nan _ _ x = true} :=
+ let (s, p) := sp in
+ exist _ (B754_nan 53 1024 s (quiet_nan_64_payload p) (quiet_nan_64_proof p)) (eq_refl true).
+
+Definition default_nan_64 := quiet_nan_64 Archi.default_nan_64.
+
+Definition quiet_nan_32_payload (p: positive) :=
+ Z.to_pos (P_mod_two_p (Pos.lor p ((iter_nat xO 22 1%positive))) 23%nat).
+
+Lemma quiet_nan_32_proof: forall p, nan_pl 24 (quiet_nan_32_payload p) = true.
+Proof. intros; apply normalized_nan; auto; omega. Qed.
+
+Definition quiet_nan_32 (sp: bool * positive) : {x :float32 | is_nan _ _ x = true} :=
+ let (s, p) := sp in
+ exist _ (B754_nan 24 128 s (quiet_nan_32_payload p) (quiet_nan_32_proof p)) (eq_refl true).
+
+Definition default_nan_32 := quiet_nan_32 Archi.default_nan_32.
+
Local Notation __ := (eq_refl Datatypes.Lt).
-Local Hint Extern 1 (Prec_gt_0 _) => exact (eq_refl Datatypes.Lt).
-Local Hint Extern 1 (_ < _) => exact (eq_refl Datatypes.Lt).
+Local Hint Extern 1 (Prec_gt_0 _) => exact (eq_refl Datatypes.Lt) : core.
+Local Hint Extern 1 (_ < _) => exact (eq_refl Datatypes.Lt) : core.
(** * Double-precision FP numbers *)
@@ -118,68 +151,44 @@ Module Float.
(** The following definitions are not part of the IEEE754 standard but
apply to all architectures supported by CompCert. *)
-(** Transform a Nan payload to a quiet Nan payload. *)
-
-Lemma transform_quiet_nan_proof (p : positive) :
- nan_pl 53 p = true ->
- nan_pl 53 (Pos.lor p (iter_nat xO 51 1%positive)) = true.
-Proof.
- unfold nan_pl. intros K.
- simpl. rewrite Z.ltb_lt, digits2_log2 in *.
- change (Z.pos (Pos.lor p 2251799813685248)) with (Z.lor (Z.pos p) 2251799813685248%Z).
- rewrite Z.log2_lor by xomega.
- now apply Z.max_case.
-Qed.
-
-Definition transform_quiet_nan s p H : {x :float | is_nan _ _ x = true} :=
- exist _ (B754_nan 53 1024 s _ (transform_quiet_nan_proof p H)) (eq_refl true).
-
(** Nan payload operations for single <-> double conversions. *)
+Definition expand_nan_payload (p: positive) := Pos.shiftl_nat p 29.
+
Lemma expand_nan_proof (p : positive) :
nan_pl 24 p = true ->
- nan_pl 53 (Pos.shiftl_nat p 29) = true.
+ nan_pl 53 (expand_nan_payload p) = true.
Proof.
- unfold nan_pl. intros K.
+ unfold nan_pl, expand_nan_payload. intros K.
rewrite Z.ltb_lt in *.
unfold Pos.shiftl_nat, nat_rect, Digits.digits2_pos.
fold (Digits.digits2_pos p).
zify; omega.
Qed.
-Definition expand_nan s p H : {x | is_nan 53 1024 x = true} :=
- exist _ (B754_nan 53 1024 s _ (expand_nan_proof p H)) (eq_refl true).
+Definition expand_nan s p H : {x | is_nan _ _ x = true} :=
+ exist _ (B754_nan 53 1024 s (expand_nan_payload p) (expand_nan_proof p H)) (eq_refl true).
Definition of_single_nan (f : float32) : { x : float | is_nan _ _ x = true } :=
match f with
| B754_nan s p H =>
if Archi.float_of_single_preserves_sNaN
then expand_nan s p H
- else transform_quiet_nan s _ (expand_nan_proof p H)
- | _ => Archi.default_nan_64
+ else quiet_nan_64 (s, expand_nan_payload p)
+ | _ => default_nan_64
end.
-Lemma reduce_nan_proof (p : positive) :
- nan_pl 53 p = true ->
- nan_pl 24 (Pos.shiftr_nat p 29) = true.
-Proof.
- unfold nan_pl. intros K.
- rewrite Z.ltb_lt in *.
- unfold Pos.shiftr_nat, nat_rect.
- assert (H : forall x, Digits.digits2_pos (Pos.div2 x) = (Digits.digits2_pos x - 1)%positive)
- by (destruct x; simpl; auto; rewrite Pplus_one_succ_r, Pos.add_sub; auto).
- rewrite !H, !Pos2Z.inj_sub_max.
- repeat (apply Z.max_lub_lt; [reflexivity |apply Z.lt_sub_lt_add_l]).
- exact K.
-Qed.
-
-Definition reduce_nan s p H : {x : float32 | is_nan _ _ x = true} :=
- exist _ (B754_nan 24 128 s _ (reduce_nan_proof p H)) (eq_refl true).
+Definition reduce_nan_payload (p: positive) :=
+ (* The [quiet_nan_64_payload p] before the right shift is redundant with
+ the [quiet_nan_32_payload p] performed after, in [to_single_nan].
+ However the former ensures that the result of the right shift is
+ not 0 and therefore representable as a positive. *)
+ Pos.shiftr_nat (quiet_nan_64_payload p) 29.
Definition to_single_nan (f : float) : { x : float32 | is_nan _ _ x = true } :=
match f with
- | B754_nan s p H => reduce_nan s _ (transform_quiet_nan_proof p H)
- | _ => Archi.default_nan_32
+ | B754_nan s p H => quiet_nan_32 (s, reduce_nan_payload p)
+ | _ => default_nan_32
end.
(** NaN payload operations for opposite and absolute value. *)
@@ -187,33 +196,62 @@ Definition to_single_nan (f : float) : { x : float32 | is_nan _ _ x = true } :=
Definition neg_nan (f : float) : { x : float | is_nan _ _ x = true } :=
match f with
| B754_nan s p H => exist _ (B754_nan 53 1024 (negb s) p H) (eq_refl true)
- | _ => Archi.default_nan_64
+ | _ => default_nan_64
end.
Definition abs_nan (f : float) : { x : float | is_nan _ _ x = true } :=
match f with
| B754_nan s p H => exist _ (B754_nan 53 1024 false p H) (eq_refl true)
- | _ => Archi.default_nan_64
+ | _ => default_nan_64
end.
-(** The NaN payload operations for two-argument arithmetic operations
- are not part of the IEEE754 standard, but all architectures of
- Compcert share a similar NaN behavior, parameterized by:
-- a "default" payload which occurs when an operation generates a NaN from
- non-NaN arguments;
-- a choice function determining which of the payload arguments to choose,
- when an operation is given two NaN arguments. *)
+(** When an arithmetic operation returns a NaN, the sign and payload
+ of this NaN are not fully specified by the IEEE standard, and vary
+ among the architectures supported by CompCert. However, the following
+ behavior applies to all the supported architectures: the payload is either
+- a default payload, independent of the arguments, or
+- the payload of one of the NaN arguments, if any.
+
+For each supported architecture, the functions [Archi.choose_nan_64]
+and [Archi.choose_nan_32] determine the payload of the result as a
+function of the payloads of the NaN arguments.
+
+Additionally, signaling NaNs are converted to quiet NaNs, as required by the standard.
+*)
+
+Definition cons_pl (x: float) (l: list (bool * positive)) :=
+ match x with B754_nan s p _ => (s, p) :: l | _ => l end.
-Definition binop_nan (x y : float) : {x : float | is_nan 53 1024 x = true} :=
- if Archi.fpu_returns_default_qNaN then Archi.default_nan_64 else
+Definition unop_nan (x: float) : {x : float | is_nan _ _ x = true} :=
+ quiet_nan_64 (Archi.choose_nan_64 (cons_pl x [])).
+
+Definition binop_nan (x y: float) : {x : float | is_nan _ _ x = true} :=
+ quiet_nan_64 (Archi.choose_nan_64 (cons_pl x (cons_pl y []))).
+
+(** For fused multiply-add, the order in which arguments are examined
+ to select a NaN payload varies across platforms. E.g. in [fma x y z],
+ x86 considers [x] first, then [y], then [z], while ARM considers [z] first,
+ then [x], then [y]. The corresponding permutation is defined
+ for each target, as function [Archi.fma_order]. *)
+
+Definition fma_nan_1 (x y z: float) : {x : float | is_nan _ _ x = true} :=
+ let '(a, b, c) := Archi.fma_order x y z in
+ quiet_nan_64 (Archi.choose_nan_64 (cons_pl a (cons_pl b (cons_pl c [])))).
+
+(** One last wrinkle for fused multiply-add: [fma zero infinity nan]
+ can return either the quiesced [nan], or the default NaN arising out
+ of the invalid operation [zero * infinity]. Of our target platforms,
+ only ARM honors the latter case. The choice between the default NaN
+ and [nan] is done as in the case of two-argument arithmetic operations. *)
+
+Definition fma_nan (x y z: float) : {x : float | is_nan _ _ x = true} :=
match x, y with
- | B754_nan s1 pl1 H1, B754_nan s2 pl2 H2 =>
- if Archi.choose_binop_pl_64 pl1 pl2
- then transform_quiet_nan s2 pl2 H2
- else transform_quiet_nan s1 pl1 H1
- | B754_nan s1 pl1 H1, _ => transform_quiet_nan s1 pl1 H1
- | _, B754_nan s2 pl2 H2 => transform_quiet_nan s2 pl2 H2
- | _, _ => Archi.default_nan_64
+ | B754_infinity _, B754_zero _ | B754_zero _, B754_infinity _ =>
+ if Archi.fma_invalid_mul_is_nan
+ then quiet_nan_64 (Archi.choose_nan_64 (Archi.default_nan_64 :: cons_pl z []))
+ else fma_nan_1 x y z
+ | _, _ =>
+ fma_nan_1 x y z
end.
(** ** Operations over double-precision floats *)
@@ -226,6 +264,8 @@ Definition eq_dec: forall (f1 f2: float), {f1 = f2} + {f1 <> f2} := Beq_dec _ _.
Definition neg: float -> float := Bopp _ _ neg_nan. (**r opposite (change sign) *)
Definition abs: float -> float := Babs _ _ abs_nan. (**r absolute value (set sign to [+]) *)
+Definition sqrt: float -> float :=
+ Bsqrt 53 1024 __ __ unop_nan mode_NE. (**r square root *)
Definition add: float -> float -> float :=
Bplus 53 1024 __ __ binop_nan mode_NE. (**r addition *)
Definition sub: float -> float -> float :=
@@ -234,6 +274,8 @@ Definition mul: float -> float -> float :=
Bmult 53 1024 __ __ binop_nan mode_NE. (**r multiplication *)
Definition div: float -> float -> float :=
Bdiv 53 1024 __ __ binop_nan mode_NE. (**r division *)
+Definition fma: float -> float -> float -> float :=
+ Bfma 53 1024 __ __ fma_nan mode_NE. (**r fused multiply-add [x * y + z] *)
Definition compare (f1 f2: float) : option Datatypes.comparison := (**r general comparison *)
Bcompare 53 1024 f1 f2.
Definition cmp (c:comparison) (f1 f2: float) : bool := (**r Boolean comparison *)
@@ -301,19 +343,15 @@ Ltac smart_omega :=
Theorem add_commut:
forall x y, is_nan _ _ x = false \/ is_nan _ _ y = false -> add x y = add y x.
Proof.
- intros. apply Bplus_commut. unfold binop_nan.
- destruct Archi.fpu_returns_default_qNaN. easy.
- destruct x, y; try reflexivity.
- now destruct H.
+ intros. apply Bplus_commut.
+ destruct x, y; try reflexivity; now destruct H.
Qed.
Theorem mul_commut:
forall x y, is_nan _ _ x = false \/ is_nan _ _ y = false -> mul x y = mul y x.
Proof.
- intros. apply Bmult_commut. unfold binop_nan.
- destruct Archi.fpu_returns_default_qNaN. easy.
- destruct x, y; try reflexivity.
- now destruct H.
+ intros. apply Bmult_commut.
+ destruct x, y; try reflexivity; now destruct H.
Qed.
(** Multiplication by 2 is diagonal addition. *)
@@ -323,9 +361,8 @@ Theorem mul2_add:
Proof.
intros. apply Bmult2_Bplus.
intros x y Hx Hy. unfold binop_nan.
- destruct Archi.fpu_returns_default_qNaN. easy.
- destruct x as [| |sx px Nx|]; try discriminate.
- now destruct y, Archi.choose_binop_pl_64.
+ destruct x; try discriminate. simpl. rewrite Archi.choose_nan_64_idem.
+ destruct y; reflexivity || discriminate.
Qed.
(** Divisions that can be turned into multiplication by an inverse. *)
@@ -337,9 +374,8 @@ Theorem div_mul_inverse:
Proof.
intros. apply Bdiv_mult_inverse. 2: easy.
intros x0 y0 z0 Hx Hy Hz. unfold binop_nan.
- destruct Archi.fpu_returns_default_qNaN. easy.
- destruct x0 as [| |sx px Nx|]; try discriminate.
- now destruct y0, z0.
+ destruct x0; try discriminate.
+ destruct y0, z0; reflexivity || discriminate.
Qed.
(** Properties of comparisons. *)
@@ -413,6 +449,7 @@ Qed.
to emulate the former.) *)
Definition ox8000_0000 := Int.repr Int.half_modulus. (**r [0x8000_0000] *)
+Definition ox7FFF_FFFF := Int.repr Int.max_signed. (**r [0x7FFF_FFFF] *)
Theorem of_intu_of_int_1:
forall x,
@@ -443,6 +480,46 @@ Proof.
compute_this (Int.unsigned ox8000_0000); smart_omega.
Qed.
+Theorem of_intu_of_int_3:
+ forall x,
+ of_intu x = sub (of_int (Int.and x ox7FFF_FFFF)) (of_int (Int.and x ox8000_0000)).
+Proof.
+ intros.
+ set (hi := Int.and x ox8000_0000).
+ set (lo := Int.and x ox7FFF_FFFF).
+ assert (R: forall n, integer_representable 53 1024 (Int.signed n)).
+ { intros. pose proof (Int.signed_range n).
+ apply integer_representable_n; auto; smart_omega. }
+ unfold sub, of_int. rewrite BofZ_minus by auto. unfold of_intu. f_equal.
+ assert (E: Int.add hi lo = x).
+ { unfold hi, lo. rewrite Int.add_is_or.
+ - rewrite <- Int.and_or_distrib. apply Int.and_mone.
+ - rewrite Int.and_assoc. rewrite (Int.and_commut ox8000_0000). rewrite Int.and_assoc.
+ change (Int.and ox7FFF_FFFF ox8000_0000) with Int.zero. rewrite ! Int.and_zero; auto.
+ }
+ assert (RNG: 0 <= Int.unsigned lo < two_p 31).
+ { unfold lo. change ox7FFF_FFFF with (Int.repr (two_p 31 - 1)). rewrite <- Int.zero_ext_and by omega.
+ apply Int.zero_ext_range. compute_this Int.zwordsize. omega. }
+ assert (B: forall i, 0 <= i < Int.zwordsize -> Int.testbit ox8000_0000 i = if zeq i 31 then true else false).
+ { intros; unfold Int.testbit. change (Int.unsigned ox8000_0000) with (2^31).
+ destruct (zeq i 31). subst i; auto. apply Z.pow2_bits_false; auto. }
+ assert (EITHER: hi = Int.zero \/ hi = ox8000_0000).
+ { unfold hi; destruct (Int.testbit x 31) eqn:B31; [right|left];
+ Int.bit_solve; rewrite B by auto.
+ - destruct (zeq i 31). subst i; rewrite B31; auto. apply andb_false_r.
+ - destruct (zeq i 31). subst i; rewrite B31; auto. apply andb_false_r.
+ }
+ assert (SU: - Int.signed hi = Int.unsigned hi).
+ { destruct EITHER as [EQ|EQ]; rewrite EQ; reflexivity. }
+ unfold Z.sub; rewrite SU, <- E.
+ unfold Int.add; rewrite Int.unsigned_repr, Int.signed_eq_unsigned. omega.
+ - assert (Int.max_signed = two_p 31 - 1) by reflexivity. omega.
+ - assert (Int.unsigned hi = 0 \/ Int.unsigned hi = two_p 31)
+ by (destruct EITHER as [EQ|EQ]; rewrite EQ; [left|right]; reflexivity).
+ assert (Int.max_unsigned = two_p 31 + two_p 31 - 1) by reflexivity.
+ omega.
+Qed.
+
Theorem to_intu_to_int_1:
forall x n,
cmp Clt x (of_intu ox8000_0000) = true ->
@@ -918,44 +995,39 @@ End Float.
Module Float32.
-(** ** NaN payload manipulations *)
-
-Lemma transform_quiet_nan_proof (p : positive) :
- nan_pl 24 p = true ->
- nan_pl 24 (Pos.lor p (iter_nat xO 22 1%positive)) = true.
-Proof.
- unfold nan_pl. intros K.
- simpl. rewrite Z.ltb_lt, digits2_log2 in *.
- change (Z.pos (Pos.lor p 4194304)) with (Z.lor (Z.pos p) 4194304%Z).
- rewrite Z.log2_lor by xomega.
- now apply Z.max_case.
-Qed.
-
-Definition transform_quiet_nan s p H : {x : float32 | is_nan _ _ x = true} :=
- exist _ (B754_nan 24 128 s _ (transform_quiet_nan_proof p H)) (eq_refl true).
-
Definition neg_nan (f : float32) : { x : float32 | is_nan _ _ x = true } :=
match f with
| B754_nan s p H => exist _ (B754_nan 24 128 (negb s) p H) (eq_refl true)
- | _ => Archi.default_nan_32
+ | _ => default_nan_32
end.
Definition abs_nan (f : float32) : { x : float32 | is_nan _ _ x = true } :=
match f with
| B754_nan s p H => exist _ (B754_nan 24 128 false p H) (eq_refl true)
- | _ => Archi.default_nan_32
+ | _ => default_nan_32
end.
-Definition binop_nan (x y : float32) : {x : float32 | is_nan _ _ x = true} :=
- if Archi.fpu_returns_default_qNaN then Archi.default_nan_32 else
+Definition cons_pl (x: float32) (l: list (bool * positive)) :=
+ match x with B754_nan s p _ => (s, p) :: l | _ => l end.
+
+Definition unop_nan (x: float32) : {x : float32 | is_nan _ _ x = true} :=
+ quiet_nan_32 (Archi.choose_nan_32 (cons_pl x [])).
+
+Definition binop_nan (x y: float32) : {x : float32 | is_nan _ _ x = true} :=
+ quiet_nan_32 (Archi.choose_nan_32 (cons_pl x (cons_pl y []))).
+
+Definition fma_nan_1 (x y z: float32) : {x : float32 | is_nan _ _ x = true} :=
+ let '(a, b, c) := Archi.fma_order x y z in
+ quiet_nan_32 (Archi.choose_nan_32 (cons_pl a (cons_pl b (cons_pl c [])))).
+
+Definition fma_nan (x y z: float32) : {x : float32 | is_nan _ _ x = true} :=
match x, y with
- | B754_nan s1 pl1 H1, B754_nan s2 pl2 H2 =>
- if Archi.choose_binop_pl_32 pl1 pl2
- then transform_quiet_nan s2 pl2 H2
- else transform_quiet_nan s1 pl1 H1
- | B754_nan s1 pl1 H1, _ => transform_quiet_nan s1 pl1 H1
- | _, B754_nan s2 pl2 H2 => transform_quiet_nan s2 pl2 H2
- | _, _ => Archi.default_nan_32
+ | B754_infinity _, B754_zero _ | B754_zero _, B754_infinity _ =>
+ if Archi.fma_invalid_mul_is_nan
+ then quiet_nan_32 (Archi.choose_nan_32 (Archi.default_nan_32 :: cons_pl z []))
+ else fma_nan_1 x y z
+ | _, _ =>
+ fma_nan_1 x y z
end.
(** ** Operations over single-precision floats *)
@@ -968,6 +1040,8 @@ Definition eq_dec: forall (f1 f2: float32), {f1 = f2} + {f1 <> f2} := Beq_dec _
Definition neg: float32 -> float32 := Bopp _ _ neg_nan. (**r opposite (change sign) *)
Definition abs: float32 -> float32 := Babs _ _ abs_nan. (**r absolute value (set sign to [+]) *)
+Definition sqrt: float32 -> float32 :=
+ Bsqrt 24 128 __ __ unop_nan mode_NE. (**r square root *)
Definition add: float32 -> float32 -> float32 :=
Bplus 24 128 __ __ binop_nan mode_NE. (**r addition *)
Definition sub: float32 -> float32 -> float32 :=
@@ -976,6 +1050,8 @@ Definition mul: float32 -> float32 -> float32 :=
Bmult 24 128 __ __ binop_nan mode_NE. (**r multiplication *)
Definition div: float32 -> float32 -> float32 :=
Bdiv 24 128 __ __ binop_nan mode_NE. (**r division *)
+Definition fma: float32 -> float32 -> float32 -> float32 :=
+ Bfma 24 128 __ __ fma_nan mode_NE. (**r fused multiply-add [x * y + z] *)
Definition compare (f1 f2: float32) : option Datatypes.comparison := (**r general comparison *)
Bcompare 24 128 f1 f2.
Definition cmp (c:comparison) (f1 f2: float32) : bool := (**r comparison *)
@@ -1023,19 +1099,15 @@ Definition of_bits (b: int): float32 := b32_of_bits (Int.unsigned b).
Theorem add_commut:
forall x y, is_nan _ _ x = false \/ is_nan _ _ y = false -> add x y = add y x.
Proof.
- intros. apply Bplus_commut. unfold binop_nan.
- destruct Archi.fpu_returns_default_qNaN. easy.
- destruct x, y; try reflexivity.
- now destruct H.
+ intros. apply Bplus_commut.
+ destruct x, y; try reflexivity; now destruct H.
Qed.
Theorem mul_commut:
forall x y, is_nan _ _ x = false \/ is_nan _ _ y = false -> mul x y = mul y x.
Proof.
- intros. apply Bmult_commut. unfold binop_nan.
- destruct Archi.fpu_returns_default_qNaN. easy.
- destruct x, y; try reflexivity.
- now destruct H.
+ intros. apply Bmult_commut.
+ destruct x, y; try reflexivity; now destruct H.
Qed.
(** Multiplication by 2 is diagonal addition. *)
@@ -1045,9 +1117,8 @@ Theorem mul2_add:
Proof.
intros. apply Bmult2_Bplus.
intros x y Hx Hy. unfold binop_nan.
- destruct Archi.fpu_returns_default_qNaN. easy.
- destruct x as [| |sx px Nx|]; try discriminate.
- now destruct y, Archi.choose_binop_pl_32.
+ destruct x; try discriminate. simpl. rewrite Archi.choose_nan_32_idem.
+ destruct y; reflexivity || discriminate.
Qed.
(** Divisions that can be turned into multiplication by an inverse. *)
@@ -1059,9 +1130,8 @@ Theorem div_mul_inverse:
Proof.
intros. apply Bdiv_mult_inverse. 2: easy.
intros x0 y0 z0 Hx Hy Hz. unfold binop_nan.
- destruct Archi.fpu_returns_default_qNaN. easy.
- destruct x0 as [| |sx px Nx|]; try discriminate.
- now destruct y0, z0.
+ destruct x0; try discriminate.
+ destruct y0, z0; reflexivity || discriminate.
Qed.
(** Properties of comparisons. *)
@@ -1218,7 +1288,7 @@ Proof.
assert (C: m / 2^p = if zeq (n mod 2^p) 0 then 0 else 1).
{ unfold m. destruct (zeq (n mod 2^p) 0).
rewrite e. apply Z.div_small. omega.
- eapply Zdiv_unique with (n mod 2^p - 1). ring. omega. }
+ eapply Coqlib.Zdiv_unique with (n mod 2^p - 1). ring. omega. }
assert (D: Z.testbit m p = if zeq (n mod 2^p) 0 then false else true).
{ destruct (zeq (n mod 2^p) 0).
apply Z.testbit_false; auto. rewrite C; auto.
diff --git a/lib/Maps.v b/lib/Maps.v
index cfb866ba..9e44a7fe 100644
--- a/lib/Maps.v
+++ b/lib/Maps.v
@@ -190,7 +190,7 @@ Module PTree <: TREE.
| Leaf : tree A
| Node : tree A -> option A -> tree A -> tree A.
- Arguments Leaf [A].
+ Arguments Leaf {A}.
Arguments Node [A].
Scheme tree_ind := Induction for tree Sort Prop.
diff --git a/mppa_k1c/Archi.v b/mppa_k1c/Archi.v
index 96571841..cdcf58c3 100644
--- a/mppa_k1c/Archi.v
+++ b/mppa_k1c/Archi.v
@@ -16,7 +16,8 @@
(** Architecture-dependent parameters for MPPA K1c. Mostly copied from the Risc-V backend *)
-Require Import ZArith.
+Require Import ZArith List.
+(*From Flocq*)
Require Import Binary Bits.
Definition ptr64 := true.
@@ -33,6 +34,8 @@ Proof.
unfold splitlong. destruct ptr64; simpl; congruence.
Qed.
+(** THIS IS NOT CHECKED ! NONE OF THIS ! *)
+
(** Section 7.3: "Except when otherwise stated, if the result of a
floating-point operation is NaN, it is the canonical NaN. The
canonical NaN has a positive sign and all significand bits clear
@@ -40,26 +43,36 @@ Qed.
We need to extend the [choose_binop_pl] functions to account for
this case. *)
-Definition default_nan_64 : { x : binary64 | is_nan _ _ x = true } :=
- exist _ (B754_nan 53 1024 true (iter_nat 51 _ xO xH) (eq_refl true)) (eq_refl true).
+Definition default_nan_64 := (false, iter_nat 51 _ xO xH).
+Definition default_nan_32 := (false, iter_nat 22 _ xO xH).
-Definition choose_binop_pl_64 (pl1 pl2 : positive) :=
- false. (**r always choose first NaN *)
+(* Always choose the first NaN argument, if any *)
-Definition default_nan_32 : { x : binary32 | is_nan _ _ x = true } :=
- exist _ (B754_nan 24 128 true (iter_nat 22 _ xO xH) (eq_refl true)) (eq_refl true).
+Definition choose_nan_64 (l: list (bool * positive)) : bool * positive :=
+ match l with nil => default_nan_64 | n :: _ => n end.
-Definition choose_binop_pl_32 (pl1 pl2 : positive) :=
- false. (**r always choose first NaN *)
+Definition choose_nan_32 (l: list (bool * positive)) : bool * positive :=
+ match l with nil => default_nan_32 | n :: _ => n end.
-(* TODO check *)
Definition fpu_returns_default_qNaN := false.
+Lemma choose_nan_64_idem: forall n,
+ choose_nan_64 (n :: n :: nil) = choose_nan_64 (n :: nil).
+Proof. auto. Qed.
+
+Lemma choose_nan_32_idem: forall n,
+ choose_nan_32 (n :: n :: nil) = choose_nan_32 (n :: nil).
+Proof. auto. Qed.
+
+Definition fma_order {A: Type} (x y z: A) := (x, z, y).
+
+Definition fma_invalid_mul_is_nan := false.
Definition float_of_single_preserves_sNaN := false.
Global Opaque ptr64 big_endian splitlong
- default_nan_64 choose_binop_pl_64
- default_nan_32 choose_binop_pl_32
+ default_nan_64 choose_nan_64
+ default_nan_32 choose_nan_32
+ fma_order fma_invalid_mul_is_nan
float_of_single_preserves_sNaN.
(** Whether to generate position-independent code or not *)
diff --git a/mppa_k1c/Asm.v b/mppa_k1c/Asm.v
index a0c5e71c..f09aa99c 100644
--- a/mppa_k1c/Asm.v
+++ b/mppa_k1c/Asm.v
@@ -204,6 +204,10 @@ Inductive instruction : Type :=
| Psllw (rd rs1 rs2: ireg) (**r shift left logical word *)
| Pmaddw (rd rs1 rs2: ireg) (**r multiply-add words *)
| Pmsubw (rd rs1 rs2: ireg) (**r multiply-add words *)
+ | Pfmaddfw (rd rs1 rs2: ireg) (**r float fused multiply-add words *)
+ | Pfmsubfw (rd rs1 rs2: ireg) (**r float fused multiply-subtract words *)
+ | Pfmaddfl (rd rs1 rs2: ireg) (**r float fused multiply-add longs *)
+ | Pfmsubfl (rd rs1 rs2: ireg) (**r float fused multiply-subtract longs *)
| Paddl (rd rs1 rs2: ireg) (**r add long *)
| Paddxl (shift : shift1_4) (rd rs1 rs2: ireg) (**r add long shift *)
@@ -231,7 +235,12 @@ Inductive instruction : Type :=
| Pfsbfw (rd rs1 rs2: ireg) (**r Float sub word *)
| Pfmuld (rd rs1 rs2: ireg) (**r Float mul double *)
| Pfmulw (rd rs1 rs2: ireg) (**r Float mul word *)
-
+ | Pfmind (rd rs1 rs2: ireg) (**r Float min double *)
+ | Pfminw (rd rs1 rs2: ireg) (**r Float min word *)
+ | Pfmaxd (rd rs1 rs2: ireg) (**r Float max double *)
+ | Pfmaxw (rd rs1 rs2: ireg) (**r Float max word *)
+ | Pfinvw (rd rs1: ireg) (**r Float invert word *)
+
(** Arith RRI32 *)
| Pcompiw (it: itest) (rd rs: ireg) (imm: int) (**r comparison imm word *)
@@ -327,6 +336,7 @@ Definition basic_to_instruction (b: basic) :=
| PArithRR Asmvliw.Pfabsw rd rs => Pfabsw rd rs
| PArithRR Asmvliw.Pfnegd rd rs => Pfnegd rd rs
| PArithRR Asmvliw.Pfnegw rd rs => Pfnegw rd rs
+ | PArithRR Asmvliw.Pfinvw rd rs => Pfinvw rd rs
| PArithRR Asmvliw.Pfnarrowdw rd rs => Pfnarrowdw rd rs
| PArithRR Asmvliw.Pfwidenlwd rd rs => Pfwidenlwd rd rs
| PArithRR Asmvliw.Pfloatuwrnsz rd rs => Pfloatuwrnsz rd rs
@@ -399,6 +409,10 @@ Definition basic_to_instruction (b: basic) :=
| PArithRRR Asmvliw.Pfsbfw rd rs1 rs2 => Pfsbfw rd rs1 rs2
| PArithRRR Asmvliw.Pfmuld rd rs1 rs2 => Pfmuld rd rs1 rs2
| PArithRRR Asmvliw.Pfmulw rd rs1 rs2 => Pfmulw rd rs1 rs2
+ | PArithRRR Asmvliw.Pfmind rd rs1 rs2 => Pfmind rd rs1 rs2
+ | PArithRRR Asmvliw.Pfminw rd rs1 rs2 => Pfminw rd rs1 rs2
+ | PArithRRR Asmvliw.Pfmaxd rd rs1 rs2 => Pfmaxd rd rs1 rs2
+ | PArithRRR Asmvliw.Pfmaxw rd rs1 rs2 => Pfmaxw rd rs1 rs2
(* RRI32 *)
| PArithRRI32 (Asmvliw.Pcompiw it) rd rs imm => Pcompiw it rd rs imm
@@ -446,6 +460,10 @@ Definition basic_to_instruction (b: basic) :=
| PArithARRR Asmvliw.Pmaddl rd rs1 rs2 => Pmaddl rd rs1 rs2
| PArithARRR Asmvliw.Pmsubw rd rs1 rs2 => Pmsubw rd rs1 rs2
| PArithARRR Asmvliw.Pmsubl rd rs1 rs2 => Pmsubl rd rs1 rs2
+ | PArithARRR Asmvliw.Pfmaddfw rd rs1 rs2 => Pfmaddfw rd rs1 rs2
+ | PArithARRR Asmvliw.Pfmaddfl rd rs1 rs2 => Pfmaddfl rd rs1 rs2
+ | PArithARRR Asmvliw.Pfmsubfw rd rs1 rs2 => Pfmsubfw rd rs1 rs2
+ | PArithARRR Asmvliw.Pfmsubfl rd rs1 rs2 => Pfmsubfl rd rs1 rs2
| PArithARRR (Asmvliw.Pcmove cond) rd rs1 rs2=> Pcmove cond rd rs1 rs2
| PArithARRR (Asmvliw.Pcmoveu cond) rd rs1 rs2=> Pcmoveu cond rd rs1 rs2
diff --git a/mppa_k1c/Asmblockdeps.v b/mppa_k1c/Asmblockdeps.v
index 2d144bb6..c4c1bbf1 100644
--- a/mppa_k1c/Asmblockdeps.v
+++ b/mppa_k1c/Asmblockdeps.v
@@ -1338,6 +1338,7 @@ Definition string_of_name_rr (n: arith_name_rr): pstring :=
| Pfabsw => "Pfabsw"
| Pfnegd => "Pfnegd"
| Pfnegw => "Pfnegw"
+ | Pfinvw => "Pfinvw"
| Pfnarrowdw => "Pfnarrowdw"
| Pfwidenlwd => "Pfwidenlwd"
| Pfloatwrnsz => "Pfloatwrnsz"
@@ -1418,6 +1419,10 @@ Definition string_of_name_rrr (n: arith_name_rrr): pstring :=
| Pfsbfw => "Pfsbfw"
| Pfmuld => "Pfmuld"
| Pfmulw => "Pfmulw"
+ | Pfmind => "Pfmind"
+ | Pfminw => "Pfminw"
+ | Pfmaxd => "Pfmaxd"
+ | Pfmaxw => "Pfmaxw"
end.
Definition string_of_name_rri32 (n: arith_name_rri32): pstring :=
@@ -1473,6 +1478,10 @@ Definition string_of_name_arrr (n: arith_name_arrr): pstring :=
| Pmsubl => "Pmsubl"
| Pcmove _ => "Pcmove"
| Pcmoveu _ => "Pcmoveu"
+ | Pfmaddfw => "Pfmaddfw"
+ | Pfmaddfl => "Pfmaddfl"
+ | Pfmsubfw => "Pfmsubfw"
+ | Pfmsubfl => "Pfmsubfl"
end.
Definition string_of_name_arr (n: arith_name_arr): pstring :=
diff --git a/mppa_k1c/Asmblockgen.v b/mppa_k1c/Asmblockgen.v
index 7e415c2a..ade84e7b 100644
--- a/mppa_k1c/Asmblockgen.v
+++ b/mppa_k1c/Asmblockgen.v
@@ -740,12 +740,52 @@ Definition transl_op
| Omulfs, a1 :: a2 :: nil =>
do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
OK (Pfmulw rd rs1 rs2 ::i k)
+ | Ominf, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfmind rd rs1 rs2 ::i k)
+ | Ominfs, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfminw rd rs1 rs2 ::i k)
+ | Omaxf, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfmaxd rd rs1 rs2 ::i k)
+ | Omaxfs, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfmaxw rd rs1 rs2 ::i k)
| Onegf, a1 :: nil =>
do rd <- freg_of res; do rs <- freg_of a1;
OK (Pfnegd rd rs ::i k)
| Onegfs, a1 :: nil =>
do rd <- freg_of res; do rs <- freg_of a1;
OK (Pfnegw rd rs ::i k)
+ | Oinvfs, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfinvw rd rs ::i k)
+
+ | Ofmaddf, a1 :: a2 :: a3 :: nil =>
+ assertion (mreg_eq a1 res);
+ do rs1 <- freg_of a1;
+ do rs2 <- freg_of a2;
+ do rs3 <- freg_of a3;
+ OK (Pfmaddfl rs1 rs2 rs3 ::i k)
+ | Ofmaddfs, a1 :: a2 :: a3 :: nil =>
+ assertion (mreg_eq a1 res);
+ do rs1 <- freg_of a1;
+ do rs2 <- freg_of a2;
+ do rs3 <- freg_of a3;
+ OK (Pfmaddfw rs1 rs2 rs3 ::i k)
+ | Ofmsubf, a1 :: a2 :: a3 :: nil =>
+ assertion (mreg_eq a1 res);
+ do rs1 <- freg_of a1;
+ do rs2 <- freg_of a2;
+ do rs3 <- freg_of a3;
+ OK (Pfmsubfl rs1 rs2 rs3 ::i k)
+ | Ofmsubfs, a1 :: a2 :: a3 :: nil =>
+ assertion (mreg_eq a1 res);
+ do rs1 <- freg_of a1;
+ do rs2 <- freg_of a2;
+ do rs3 <- freg_of a3;
+ OK (Pfmsubfw rs1 rs2 rs3 ::i k)
| Osingleofint, a1 :: nil =>
do rd <- freg_of res; do rs <- ireg_of a1;
diff --git a/mppa_k1c/Asmexpand.ml b/mppa_k1c/Asmexpand.ml
index 20d27951..1e5149fd 100644
--- a/mppa_k1c/Asmexpand.ml
+++ b/mppa_k1c/Asmexpand.ml
@@ -467,11 +467,11 @@ let expand_builtin_inline name args res = let open Asmvliw in
emit (Pdzerol addr)
(*| "__builtin_k1_afaddd", [BA(IR addr); BA (IR incr_res)], BR(IR res) ->
(if res <> incr_res
- then (emit (Pmv(res, incr_res)); emit Psemi));
+ then (emit (Asm.Pmv(res, incr_res)); emit Psemi));
emit (Pafaddd(addr, res))
| "__builtin_k1_afaddw", [BA(IR addr); BA (IR incr_res)], BR(IR res) ->
(if res <> incr_res
- then (emit (Pmv(res, incr_res)); emit Psemi));
+ then (emit (Asm.Pmv(res, incr_res)); emit Psemi));
emit (Pafaddw(addr, res)) *) (* see #157 *)
| "__builtin_alclrd", [BA(IR addr)], BR(IR res) ->
emit (Palclrd(res, addr))
diff --git a/mppa_k1c/Asmvliw.v b/mppa_k1c/Asmvliw.v
index 72584d2a..54654abb 100644
--- a/mppa_k1c/Asmvliw.v
+++ b/mppa_k1c/Asmvliw.v
@@ -375,6 +375,7 @@ Inductive arith_name_rr : Type :=
| Pfabsw (**r float absolute word *)
| Pfnegd (**r float negate double *)
| Pfnegw (**r float negate word *)
+ | Pfinvw (**r float invert word *)
| Pfnarrowdw (**r float narrow 64 -> 32 bits *)
| Pfwidenlwd (**r Floating Point widen from 32 bits to 64 bits *)
| Pfloatwrnsz (**r Floating Point conversion from integer (int -> SINGLE) *)
@@ -453,6 +454,10 @@ Inductive arith_name_rrr : Type :=
| Pfsbfw (**r float sub word *)
| Pfmuld (**r float multiply double *)
| Pfmulw (**r float multiply word *)
+ | Pfmind (**r float min double *)
+ | Pfminw (**r float min word *)
+ | Pfmaxd (**r float max double *)
+ | Pfmaxw (**r float max word *)
.
Inductive arith_name_rri32 : Type :=
@@ -506,6 +511,10 @@ Inductive arith_name_arrr : Type :=
| Pmsubl (**r multiply subtract long *)
| Pcmove (bt: btest) (**r conditional move *)
| Pcmoveu (bt: btest) (**r conditional move, test on unsigned semantics *)
+ | Pfmaddfw (**r float fused multiply add word *)
+ | Pfmaddfl (**r float fused multiply add long *)
+ | Pfmsubfw (**r float fused multiply subtract word *)
+ | Pfmsubfl (**r float fused multiply subtract long *)
.
Inductive arith_name_arri32 : Type :=
@@ -975,6 +984,7 @@ Definition arith_eval_rr n v :=
| Pfnegw => Val.negfs v
| Pfabsd => Val.absf v
| Pfabsw => Val.absfs v
+ | Pfinvw => ExtValues.invfs v
| Pfnarrowdw => Val.singleoffloat v
| Pfwidenlwd => Val.floatofsingle v
| Pfloatwrnsz => match Val.singleofint v with Some f => f | _ => Vundef end
@@ -1055,6 +1065,11 @@ Definition arith_eval_rrr n v1 v2 :=
| Pfmuld => Val.mulf v1 v2
| Pfmulw => Val.mulfs v1 v2
+ | Pfmind => ExtValues.minf v1 v2
+ | Pfminw => ExtValues.minfs v1 v2
+ | Pfmaxd => ExtValues.maxf v1 v2
+ | Pfmaxw => ExtValues.maxfs v1 v2
+
| Paddxw shift => ExtValues.addx (int_of_shift1_4 shift) v1 v2
| Paddxl shift => ExtValues.addxl (int_of_shift1_4 shift) v1 v2
@@ -1149,6 +1164,10 @@ Definition arith_eval_arrr n v1 v2 v3 :=
| Pmsubl => Val.subl v1 (Val.mull v2 v3)
| Pcmove bt => cmove bt v1 v2 v3
| Pcmoveu bt => cmoveu bt v1 v2 v3
+ | Pfmaddfw => ExtValues.fmaddfs v1 v2 v3
+ | Pfmaddfl => ExtValues.fmaddf v1 v2 v3
+ | Pfmsubfw => ExtValues.fmsubfs v1 v2 v3
+ | Pfmsubfl => ExtValues.fmsubf v1 v2 v3
end.
Definition arith_eval_arr n v1 v2 :=
diff --git a/mppa_k1c/Builtins1.v b/mppa_k1c/Builtins1.v
new file mode 100644
index 00000000..6186961f
--- /dev/null
+++ b/mppa_k1c/Builtins1.v
@@ -0,0 +1,66 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, Collège de France and Inria Paris *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the GNU General Public License as published by *)
+(* the Free Software Foundation, either version 2 of the License, or *)
+(* (at your option) any later version. This file is also distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Platform-specific built-in functions *)
+
+Require Import String Coqlib.
+Require Import AST Integers Floats Values ExtFloats.
+Require Import Builtins0.
+
+Inductive platform_builtin : Type :=
+| BI_fmin
+| BI_fmax
+| BI_fminf
+| BI_fmaxf
+| BI_fabsf
+| BI_fma
+| BI_fmaf.
+
+Local Open Scope string_scope.
+
+Definition platform_builtin_table : list (string * platform_builtin) :=
+ ("__builtin_fmin", BI_fmin)
+ :: ("__builtin_fmax", BI_fmax)
+ :: ("__builtin_fminf", BI_fminf)
+ :: ("__builtin_fmaxf", BI_fmaxf)
+ :: ("__builtin_fabsf", BI_fabsf)
+ :: ("__builtin_fma", BI_fma)
+ :: ("__builtin_fmaf", BI_fmaf)
+ :: nil.
+
+Definition platform_builtin_sig (b: platform_builtin) : signature :=
+ match b with
+ | BI_fmin | BI_fmax =>
+ mksignature (Tfloat :: Tfloat :: nil) (Some Tfloat) cc_default
+ | BI_fminf | BI_fmaxf =>
+ mksignature (Tsingle :: Tsingle :: nil) (Some Tsingle) cc_default
+ | BI_fabsf =>
+ mksignature (Tsingle :: nil) (Some Tsingle) cc_default
+ | BI_fma =>
+ mksignature (Tfloat :: Tfloat :: Tfloat :: nil) (Some Tfloat) cc_default
+ | BI_fmaf =>
+ mksignature (Tsingle :: Tsingle :: Tsingle :: nil) (Some Tsingle) cc_default
+ end.
+
+Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (proj_sig_res (platform_builtin_sig b)) :=
+ match b with
+ | BI_fmin => mkbuiltin_n2t Tfloat Tfloat Tfloat ExtFloat.min
+ | BI_fmax => mkbuiltin_n2t Tfloat Tfloat Tfloat ExtFloat.max
+ | BI_fminf => mkbuiltin_n2t Tsingle Tsingle Tsingle ExtFloat32.min
+ | BI_fmaxf => mkbuiltin_n2t Tsingle Tsingle Tsingle ExtFloat32.max
+ | BI_fabsf => mkbuiltin_n1t Tsingle Tsingle Float32.abs
+ | BI_fma => mkbuiltin_n3t Tfloat Tfloat Tfloat Tfloat Float.fma
+ | BI_fmaf => mkbuiltin_n3t Tsingle Tsingle Tsingle Tsingle Float32.fma
+ end.
diff --git a/mppa_k1c/CBuiltins.ml b/mppa_k1c/CBuiltins.ml
index 5fb69f62..235496c5 100644
--- a/mppa_k1c/CBuiltins.ml
+++ b/mppa_k1c/CBuiltins.ml
@@ -18,11 +18,11 @@
open C
let builtins = {
- Builtins.typedefs = [
+ builtin_typedefs = [
"__builtin_va_list", TPtr(TVoid [], [])
];
(* The builtin list is inspired from the GCC file builtin_k1.h *)
- Builtins.functions = [ (* Some builtins are commented out because their opcode is not present (yet?) *)
+ builtin_functions = [ (* Some builtins are commented out because their opcode is not present (yet?) *)
(* BCU Instructions *)
"__builtin_k1_await", (TVoid [], [], false); (* DONE *)
"__builtin_k1_barrier", (TVoid [], [], false); (* DONE *)
@@ -113,14 +113,29 @@ let builtins = {
[TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])], false);
"__builtin_fnmsub",
(TFloat(FDouble, []),
- [TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])], false);
+ [TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])], false); *)
+ "__builtin_fabsf",
+ (TFloat(FFloat, []),
+ [TFloat(FFloat, [])], false);
"__builtin_fmax",
(TFloat(FDouble, []),
[TFloat(FDouble, []); TFloat(FDouble, [])], false);
"__builtin_fmin",
(TFloat(FDouble, []),
[TFloat(FDouble, []); TFloat(FDouble, [])], false);
-*)]
+ "__builtin_fmaxf",
+ (TFloat(FFloat, []),
+ [TFloat(FFloat, []); TFloat(FFloat, [])], false);
+ "__builtin_fminf",
+ (TFloat(FFloat, []),
+ [TFloat(FFloat, []); TFloat(FFloat, [])], false);
+ "__builtin_fma",
+ (TFloat(FDouble, []),
+ [TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])], false);
+ "__builtin_fmaf",
+ (TFloat(FFloat, []),
+ [TFloat(FFloat, []); TFloat(FFloat, []); TFloat(FFloat, [])], false);
+]
}
let va_list_type = TPtr(TVoid [], []) (* to check! *)
diff --git a/mppa_k1c/ExtFloats.v b/mppa_k1c/ExtFloats.v
new file mode 100644
index 00000000..d9b9d3a6
--- /dev/null
+++ b/mppa_k1c/ExtFloats.v
@@ -0,0 +1,39 @@
+Require Import Floats Integers ZArith.
+
+Module ExtFloat.
+(** TODO check with the actual K1c;
+ this is what happens on x86 and may be inappropriate. *)
+
+Definition min (x : float) (y : float) : float :=
+ match Float.compare x y with
+ | Some Eq | Some Lt => x
+ | Some Gt | None => y
+ end.
+
+Definition max (x : float) (y : float) : float :=
+ match Float.compare x y with
+ | Some Eq | Some Gt => x
+ | Some Lt | None => y
+ end.
+End ExtFloat.
+
+Module ExtFloat32.
+(** TODO check with the actual K1c *)
+
+Definition min (x : float32) (y : float32) : float32 :=
+ match Float32.compare x y with
+ | Some Eq | Some Lt => x
+ | Some Gt | None => y
+ end.
+
+Definition max (x : float32) (y : float32) : float32 :=
+ match Float32.compare x y with
+ | Some Eq | Some Gt => x
+ | Some Lt | None => y
+ end.
+
+Definition one := Float32.of_int (Int.repr (1%Z)).
+Definition inv (x : float32) : float32 :=
+ Float32.div one x.
+
+End ExtFloat32.
diff --git a/mppa_k1c/ExtValues.v b/mppa_k1c/ExtValues.v
index e9c62a8d..5a890f3c 100644
--- a/mppa_k1c/ExtValues.v
+++ b/mppa_k1c/ExtValues.v
@@ -1,6 +1,7 @@
Require Import Coqlib.
Require Import Integers.
Require Import Values.
+Require Import Floats ExtFloats.
Open Scope Z_scope.
@@ -689,3 +690,51 @@ Definition revsubx sh v1 v2 :=
Definition revsubxl sh v1 v2 :=
Val.subl v2 (Val.shll v1 (Vint sh)).
+
+Definition minf v1 v2 :=
+ match v1, v2 with
+ | (Vfloat f1), (Vfloat f2) => Vfloat (ExtFloat.min f1 f2)
+ | _, _ => Vundef
+ end.
+
+Definition maxf v1 v2 :=
+ match v1, v2 with
+ | (Vfloat f1), (Vfloat f2) => Vfloat (ExtFloat.max f1 f2)
+ | _, _ => Vundef
+ end.
+
+Definition minfs v1 v2 :=
+ match v1, v2 with
+ | (Vsingle f1), (Vsingle f2) => Vsingle (ExtFloat32.min f1 f2)
+ | _, _ => Vundef
+ end.
+
+Definition maxfs v1 v2 :=
+ match v1, v2 with
+ | (Vsingle f1), (Vsingle f2) => Vsingle (ExtFloat32.max f1 f2)
+ | _, _ => Vundef
+ end.
+
+Definition invfs v1 :=
+ match v1 with
+ | (Vsingle f1) => Vsingle (ExtFloat32.inv f1)
+ | _ => Vundef
+ end.
+
+Definition triple_op_float f v1 v2 v3 :=
+ match v1, v2, v3 with
+ | (Vfloat f1), (Vfloat f2), (Vfloat f3) => Vfloat (f f1 f2 f3)
+ | _, _, _ => Vundef
+ end.
+
+Definition triple_op_single f v1 v2 v3 :=
+ match v1, v2, v3 with
+ | (Vsingle f1), (Vsingle f2), (Vsingle f3) => Vsingle (f f1 f2 f3)
+ | _, _, _ => Vundef
+ end.
+
+Definition fmaddf := triple_op_float (fun f1 f2 f3 => Float.fma f2 f3 f1).
+Definition fmaddfs := triple_op_single (fun f1 f2 f3 => Float32.fma f2 f3 f1).
+
+Definition fmsubf := triple_op_float (fun f1 f2 f3 => Float.fma (Float.neg f2) f3 f1).
+Definition fmsubfs := triple_op_single (fun f1 f2 f3 => Float32.fma (Float32.neg f2) f3 f1).
diff --git a/mppa_k1c/Machregs.v b/mppa_k1c/Machregs.v
index 5a7d42b4..8098b5d1 100644
--- a/mppa_k1c/Machregs.v
+++ b/mppa_k1c/Machregs.v
@@ -213,6 +213,8 @@ Global Opaque
Definition two_address_op (op: operation) : bool :=
match op with
+ | Ofmaddf | Ofmaddfs
+ | Ofmsubf | Ofmsubfs
| Omadd | Omaddimm _
| Omaddl | Omaddlimm _
| Omsub | Omsubl
diff --git a/mppa_k1c/NeedOp.v b/mppa_k1c/NeedOp.v
index 4748f38b..d2d4d5f5 100644
--- a/mppa_k1c/NeedOp.v
+++ b/mppa_k1c/NeedOp.v
@@ -121,9 +121,12 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval :=
| Omaddlimm n => op2 (default nv)
| Omsubl => op3 (default nv)
| Onegf | Oabsf => op1 (default nv)
- | Oaddf | Osubf | Omulf | Odivf => op2 (default nv)
+ | Oaddf | Osubf | Omulf | Odivf | Ominf | Omaxf => op2 (default nv)
+ | Ofmaddf | Ofmsubf => op3 (default nv)
| Onegfs | Oabsfs => op1 (default nv)
- | Oaddfs | Osubfs | Omulfs | Odivfs => op2 (default nv)
+ | Oaddfs | Osubfs | Omulfs | Odivfs | Ominfs | Omaxfs => op2 (default nv)
+ | Oinvfs => op1 (default nv)
+ | Ofmaddfs | Ofmsubfs => op3 (default nv)
| Ofloatofsingle | Osingleoffloat => op1 (default nv)
| Ointoffloat | Ointuoffloat => op1 (default nv)
| Olongoffloat | Olonguoffloat | Ofloatoflong | Ofloatoflongu => op1 (default nv)
@@ -285,7 +288,53 @@ Remark default_idem: forall nv, default (default nv) = default nv.
Proof.
destruct nv; simpl; trivial.
Qed.
-
+
+Lemma vagree_triple_op_float :
+ forall f a b c x y z nv,
+ (vagree a x (default nv)) ->
+ (vagree b y (default nv)) ->
+ (vagree c z (default nv)) ->
+ (vagree (ExtValues.triple_op_float f a b c)
+ (ExtValues.triple_op_float f x y z) nv).
+Proof.
+ induction nv;
+ intros Hax Hby Hcz.
+ - trivial.
+ - simpl in *. destruct a; simpl; trivial.
+ destruct b; simpl; trivial.
+ destruct c; simpl; trivial.
+ - simpl in *. destruct a; simpl; trivial.
+ destruct b; simpl; trivial.
+ destruct c; simpl; trivial.
+ inv Hax. inv Hby. inv Hcz.
+ simpl.
+ constructor.
+Qed.
+
+Lemma vagree_triple_op_single :
+ forall f a b c x y z nv,
+ (vagree a x (default nv)) ->
+ (vagree b y (default nv)) ->
+ (vagree c z (default nv)) ->
+ (vagree (ExtValues.triple_op_single f a b c)
+ (ExtValues.triple_op_single f x y z) nv).
+Proof.
+ induction nv;
+ intros Hax Hby Hcz.
+ - trivial.
+ - simpl in *. destruct a; simpl; trivial.
+ destruct b; simpl; trivial.
+ destruct c; simpl; trivial.
+ - simpl in *. destruct a; simpl; trivial.
+ destruct b; simpl; trivial.
+ destruct c; simpl; trivial.
+ inv Hax. inv Hby. inv Hcz.
+ simpl.
+ constructor.
+Qed.
+
+Hint Resolve vagree_triple_op_float vagree_triple_op_single : na.
+
Lemma needs_of_operation_sound:
forall op args v nv args',
eval_operation ge (Vptr sp Ptrofs.zero) op args m1 = Some v ->
@@ -344,7 +393,10 @@ Proof.
apply mull_sound; trivial.
rewrite default_idem; trivial.
rewrite default_idem; trivial.
- (* select *)
+- apply vagree_triple_op_float; assumption.
+- apply vagree_triple_op_float; assumption.
+- apply vagree_triple_op_single; assumption.
+- apply vagree_triple_op_single; assumption.
- destruct (eval_condition0 _ _ _) as [b|] eqn:EC.
erewrite needs_of_condition0_sound by eauto.
apply select_sound; auto.
diff --git a/mppa_k1c/Op.v b/mppa_k1c/Op.v
index f3ee0577..c75a1a22 100644
--- a/mppa_k1c/Op.v
+++ b/mppa_k1c/Op.v
@@ -179,12 +179,21 @@ Inductive operation : Type :=
| Osubf (**r [rd = r1 - r2] *)
| Omulf (**r [rd = r1 * r2] *)
| Odivf (**r [rd = r1 / r2] *)
+ | Ominf
+ | Omaxf
+ | Ofmaddf
+ | Ofmsubf
| Onegfs (**r [rd = - r1] *)
| Oabsfs (**r [rd = abs(r1)] *)
| Oaddfs (**r [rd = r1 + r2] *)
| Osubfs (**r [rd = r1 - r2] *)
| Omulfs (**r [rd = r1 * r2] *)
| Odivfs (**r [rd = r1 / r2] *)
+ | Ominfs
+ | Omaxfs
+ | Oinvfs
+ | Ofmaddfs
+ | Ofmsubfs
| Osingleoffloat (**r [rd] is [r1] truncated to single-precision float *)
| Ofloatofsingle (**r [rd] is [r1] extended to double-precision float *)
(*c Conversions between int and float: *)
@@ -426,12 +435,23 @@ Definition eval_operation
| Osubf, v1::v2::nil => Some (Val.subf v1 v2)
| Omulf, v1::v2::nil => Some (Val.mulf v1 v2)
| Odivf, v1::v2::nil => Some (Val.divf v1 v2)
+ | Ominf, v1::v2::nil => Some (ExtValues.minf v1 v2)
+ | Omaxf, v1::v2::nil => Some (ExtValues.maxf v1 v2)
+ | Ofmaddf, v1::v2::v3::nil => Some (ExtValues.fmaddf v1 v2 v3)
+ | Ofmsubf, v1::v2::v3::nil => Some (ExtValues.fmsubf v1 v2 v3)
+
| Onegfs, v1::nil => Some (Val.negfs v1)
| Oabsfs, v1::nil => Some (Val.absfs v1)
| Oaddfs, v1::v2::nil => Some (Val.addfs v1 v2)
| Osubfs, v1::v2::nil => Some (Val.subfs v1 v2)
| Omulfs, v1::v2::nil => Some (Val.mulfs v1 v2)
| Odivfs, v1::v2::nil => Some (Val.divfs v1 v2)
+ | Ominfs, v1::v2::nil => Some (ExtValues.minfs v1 v2)
+ | Omaxfs, v1::v2::nil => Some (ExtValues.maxfs v1 v2)
+ | Oinvfs, v1::nil => Some (ExtValues.invfs v1)
+ | Ofmaddfs, v1::v2::v3::nil => Some (ExtValues.fmaddfs v1 v2 v3)
+ | Ofmsubfs, v1::v2::v3::nil => Some (ExtValues.fmsubfs v1 v2 v3)
+
| Osingleoffloat, v1::nil => Some (Val.singleoffloat v1)
| Ofloatofsingle, v1::nil => Some (Val.floatofsingle v1)
| Ointoffloat, v1::nil => Val.intoffloat v1
@@ -630,16 +650,25 @@ Definition type_of_operation (op: operation) : list typ * typ :=
| Onegf => (Tfloat :: nil, Tfloat)
| Oabsf => (Tfloat :: nil, Tfloat)
- | Oaddf => (Tfloat :: Tfloat :: nil, Tfloat)
- | Osubf => (Tfloat :: Tfloat :: nil, Tfloat)
- | Omulf => (Tfloat :: Tfloat :: nil, Tfloat)
- | Odivf => (Tfloat :: Tfloat :: nil, Tfloat)
+ | Oaddf
+ | Osubf
+ | Omulf
+ | Odivf
+ | Ominf
+ | Omaxf => (Tfloat :: Tfloat :: nil, Tfloat)
+ | Ofmaddf | Ofmsubf => (Tfloat :: Tfloat :: Tfloat :: nil, Tfloat)
+
| Onegfs => (Tsingle :: nil, Tsingle)
| Oabsfs => (Tsingle :: nil, Tsingle)
- | Oaddfs => (Tsingle :: Tsingle :: nil, Tsingle)
- | Osubfs => (Tsingle :: Tsingle :: nil, Tsingle)
- | Omulfs => (Tsingle :: Tsingle :: nil, Tsingle)
- | Odivfs => (Tsingle :: Tsingle :: nil, Tsingle)
+ | Oaddfs
+ | Osubfs
+ | Omulfs
+ | Odivfs
+ | Ominfs
+ | Omaxfs => (Tsingle :: Tsingle :: nil, Tsingle)
+ | Oinvfs => (Tsingle :: nil, Tsingle)
+ | Ofmaddfs | Ofmsubfs => (Tsingle :: Tsingle :: Tsingle :: nil, Tsingle)
+
| Osingleoffloat => (Tfloat :: nil, Tsingle)
| Ofloatofsingle => (Tsingle :: nil, Tfloat)
| Ointoffloat => (Tfloat :: nil, Tint)
@@ -906,6 +935,12 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
(* mulf, divf *)
- destruct v0; destruct v1...
- destruct v0; destruct v1...
+ (* minf, maxf *)
+ - destruct v0; destruct v1...
+ - destruct v0; destruct v1...
+ (* fmaddf, fmsubf *)
+ - destruct v0; destruct v1; destruct v2...
+ - destruct v0; destruct v1; destruct v2...
(* negfs, absfs *)
- destruct v0...
- destruct v0...
@@ -915,6 +950,14 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
(* mulfs, divfs *)
- destruct v0; destruct v1...
- destruct v0; destruct v1...
+ (* minfs, maxfs *)
+ - destruct v0; destruct v1...
+ - destruct v0; destruct v1...
+ (* invfs *)
+ - destruct v0...
+ (* fmaddfs, fmsubfs *)
+ - destruct v0; destruct v1; destruct v2...
+ - destruct v0; destruct v1; destruct v2...
(* singleoffloat, floatofsingle *)
- destruct v0...
- destruct v0...
@@ -1517,6 +1560,12 @@ Proof.
(* mulf, divf *)
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
+ (* minf, maxf *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; simpl; auto.
+ (* fmaddf, fmsubf *)
+ - inv H4; inv H3; inv H2; simpl; auto.
+ - inv H4; inv H3; inv H2; simpl; auto.
(* negfs, absfs *)
- inv H4; simpl; auto.
- inv H4; simpl; auto.
@@ -1526,6 +1575,14 @@ Proof.
(* mulfs, divfs *)
- inv H4; inv H2; simpl; auto.
- inv H4; inv H2; simpl; auto.
+ (* minfs, maxfs *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; simpl; auto.
+ (* invfs *)
+ - inv H4; simpl; auto.
+ (* fmaddfs, fmsubfs *)
+ - inv H4; inv H3; inv H2; simpl; auto.
+ - inv H4; inv H3; inv H2; simpl; auto.
(* singleoffloat, floatofsingle *)
- inv H4; simpl; auto.
- inv H4; simpl; auto.
diff --git a/mppa_k1c/PostpassSchedulingOracle.ml b/mppa_k1c/PostpassSchedulingOracle.ml
index 40f1d9c7..fa61d588 100644
--- a/mppa_k1c/PostpassSchedulingOracle.ml
+++ b/mppa_k1c/PostpassSchedulingOracle.ml
@@ -31,6 +31,8 @@ type real_instruction =
(* FPU *)
| Fabsd | Fabsw | Fnegw | Fnegd
| Faddd | Faddw | Fsbfd | Fsbfw | Fmuld | Fmulw
+ | Fmind | Fminw | Fmaxd | Fmaxw | Finvw
+ | Ffmaw | Ffmad | Ffmsw | Ffmsd
| Fnarrowdw | Fwidenlwd | Floatwz | Floatuwz | Floatdz | Floatudz | Fixedwz | Fixeduwz | Fixeddz | Fixedudz
| Fcompw | Fcompd
@@ -63,6 +65,7 @@ let arith_rr_real = function
| Pfabsd -> Fabsd
| Pfnegw -> Fnegw
| Pfnegd -> Fnegd
+ | Pfinvw -> Finvw
| Pfnarrowdw -> Fnarrowdw
| Pfwidenlwd -> Fwidenlwd
| Pfloatwrnsz -> Floatwz
@@ -121,6 +124,10 @@ let arith_rrr_real = function
| Pfsbfw -> Fsbfw
| Pfmuld -> Fmuld
| Pfmulw -> Fmulw
+ | Pfmind -> Fmind
+ | Pfminw -> Fminw
+ | Pfmaxd -> Fmaxd
+ | Pfmaxw -> Fmaxw
let arith_rri32_real = function
| Pcompiw it -> Compw
@@ -169,6 +176,10 @@ let arith_arr_real = function
| Pinsfl (_, _) -> Insf
let arith_arrr_real = function
+ | Pfmaddfw -> Ffmaw
+ | Pfmaddfl -> Ffmad
+ | Pfmsubfw -> Ffmsw
+ | Pfmsubfl -> Ffmsd
| Pmaddw -> Maddw
| Pmaddl -> Maddd
| Pmsubw -> Msbfw
@@ -610,9 +621,12 @@ let rec_to_usage r =
| Some E27U27L10 -> lsu_acc_y)
| Icall | Call | Cb | Igoto | Goto | Ret | Set -> bcu
| Get -> bcu_tiny_tiny_mau_xnop
- | Fnegd | Fnegw | Fabsd | Fabsw | Fwidenlwd -> alu_lite
+ | Fnegd | Fnegw | Fabsd | Fabsw | Fwidenlwd
+ | Fmind | Fmaxd | Fminw | Fmaxw -> alu_lite
| Fnarrowdw -> alu_full
- | Faddd | Faddw | Fsbfd | Fsbfw | Fmuld | Fmulw -> mau
+ | Faddd | Faddw | Fsbfd | Fsbfw | Fmuld | Fmulw | Finvw
+ | Ffmad | Ffmaw | Ffmsd | Ffmsw -> mau
+
let inst_info_to_dlatency i =
begin
@@ -632,6 +646,7 @@ let real_inst_to_latency = function
| Nandd | Nord | Nxord | Ornd | Andnd
| Addd | Andd | Compd | Ord | Sbfd | Sbfxd | Srad | Srsd | Srld | Slld | Xord | Make
| Extfs | Extfz | Insf | Fcompw | Fcompd | Cmoved | Addxw | Addxd
+ | Fmind | Fmaxd | Fminw | Fmaxw
-> 1
| Floatwz | Floatuwz | Fixeduwz | Fixedwz | Floatdz | Floatudz | Fixeddz | Fixedudz -> 4
| Mulw | Muld | Maddw | Maddd | Msbfw | Msbfd -> 2 (* FIXME - WORST CASE. If it's S10 then it's only 1 *)
@@ -641,7 +656,8 @@ let real_inst_to_latency = function
| Set -> 4 (* According to the manual should be 3, but I measured 4 *)
| Icall | Call | Cb | Igoto | Goto | Ret -> 42 (* Should not matter since it's the final instruction of the basic block *)
| Fnegd | Fnegw | Fabsd | Fabsw | Fwidenlwd | Fnarrowdw -> 1
- | Faddd | Faddw | Fsbfd | Fsbfw | Fmuld | Fmulw -> 4
+ | Faddd | Faddw | Fsbfd | Fsbfw | Fmuld | Fmulw | Finvw
+ | Ffmaw | Ffmad | Ffmsw | Ffmsd -> 4
let rec empty_inter la = function
| [] -> true
diff --git a/mppa_k1c/SelectOp.vp b/mppa_k1c/SelectOp.vp
index 3df0c682..ec3985c5 100644
--- a/mppa_k1c/SelectOp.vp
+++ b/mppa_k1c/SelectOp.vp
@@ -51,9 +51,10 @@ Require Import Floats.
Require Import Op.
Require Import CminorSel.
Require Import OpHelpers.
-Require Import ExtValues.
+Require Import ExtValues ExtFloats.
Require Import DecBoolOps.
Require Import Chunks.
+Require Import Builtins.
Require Compopts.
Local Open Scope cminorsel_scope.
@@ -668,9 +669,45 @@ Definition divf_base (e1: expr) (e2: expr) :=
(* Eop Odivf (e1 ::: e2 ::: Enil). *)
Eexternal f64_div sig_ff_f (e1 ::: e2 ::: Enil).
-Definition divfs_base (e1: expr) (e2: expr) :=
+Definition divfs_base1 (e2 : expr) :=
+ Eop Oinvfs (e2 ::: Enil).
+Definition divfs_baseX (e1 : expr) (e2 : expr) :=
(* Eop Odivf (e1 ::: e2 ::: Enil). *)
Eexternal f32_div sig_ss_s (e1 ::: e2 ::: Enil).
+
+Nondetfunction divfs_base (e1: expr) :=
+ match e1 with
+ | Eop (Osingleconst f) Enil =>
+ (if Float32.eq_dec f ExtFloat32.one
+ then divfs_base1
+ else divfs_baseX e1)
+ | _ => divfs_baseX e1
+ end.
+
+Nondetfunction gen_fma args :=
+ match args with
+ | (Eop Onegf (e1:::Enil)):::e2:::e3:::Enil => Some (Eop Ofmsubf (e3:::e1:::e2:::Enil))
+ | e1:::e2:::e3:::Enil => Some (Eop Ofmaddf (e3:::e1:::e2:::Enil))
+ | _ => None
+ end.
+
+Nondetfunction gen_fmaf args :=
+ match args with
+ | (Eop Onegfs (e1:::Enil)):::e2:::e3:::Enil => Some (Eop Ofmsubfs (e3:::e1:::e2:::Enil))
+ | e1:::e2:::e3:::Enil => Some (Eop Ofmaddfs (e3:::e1:::e2:::Enil))
+ | _ => None
+ end.
+
+Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr :=
+ match b with
+ | BI_fmin => Some (Eop Ominf args)
+ | BI_fmax => Some (Eop Omaxf args)
+ | BI_fminf => Some (Eop Ominfs args)
+ | BI_fmaxf => Some (Eop Omaxfs args)
+ | BI_fabsf => Some (Eop Oabsfs args)
+ | BI_fma => gen_fma args
+ | BI_fmaf => gen_fmaf args
+ end.
End SELECT.
(* Local Variables: *)
diff --git a/mppa_k1c/SelectOpproof.v b/mppa_k1c/SelectOpproof.v
index 21a06857..583fb545 100644
--- a/mppa_k1c/SelectOpproof.v
+++ b/mppa_k1c/SelectOpproof.v
@@ -17,6 +17,7 @@
(** Correctness of instruction selection for operators *)
+Require Import Builtins.
Require Import Coqlib.
Require Import Maps.
Require Import AST.
@@ -29,6 +30,7 @@ Require Import Globalenvs.
Require Import Cminor.
Require Import Op.
Require Import CminorSel.
+Require Import Builtins1.
Require Import SelectOp.
Require Import Events.
Require Import OpHelpers.
@@ -1620,6 +1622,29 @@ Proof.
econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto.
Qed.
+
+Lemma eval_divfs_base1:
+ forall le a b x y,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ exists v, eval_expr ge sp e m le (divfs_base1 b) v /\ Val.lessdef (ExtValues.invfs y) v.
+Proof.
+ intros; unfold divfs_base1.
+ econstructor; split.
+ repeat (try econstructor; try eassumption).
+ trivial.
+Qed.
+
+Lemma eval_divfs_baseX:
+ forall le a b x y,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ exists v, eval_expr ge sp e m le (divfs_baseX a b) v /\ Val.lessdef (Val.divfs x y) v.
+Proof.
+ intros; unfold divfs_base.
+ econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto.
+Qed.
+
Theorem eval_divfs_base:
forall le a b x y,
eval_expr ge sp e m le a x ->
@@ -1627,6 +1652,82 @@ Theorem eval_divfs_base:
exists v, eval_expr ge sp e m le (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v.
Proof.
intros; unfold divfs_base.
- econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto.
+ destruct (divfs_base_match _).
+ - destruct (Float32.eq_dec _ _).
+ + exists (Val.divfs x y).
+ split; trivial. repeat (try econstructor; try eassumption).
+ simpl. InvEval. reflexivity.
+ + apply eval_divfs_baseX; assumption.
+ - apply eval_divfs_baseX; assumption.
Qed.
+
+(** Platform-specific known builtins *)
+
+Lemma eval_fma:
+ forall al a vl v le,
+ gen_fma al = Some a ->
+ eval_exprlist ge sp e m le al vl ->
+ platform_builtin_sem BI_fma vl = Some v ->
+ exists v', eval_expr ge sp e m le a v' /\ Val.lessdef v v'.
+Proof.
+ unfold gen_fma.
+ intros until le.
+ intro Heval.
+ destruct (gen_fma_match _) in *; try discriminate.
+ all: inversion Heval; subst a; clear Heval; intro; InvEval.
+ - subst v1.
+ TrivialExists.
+ destruct v0; simpl; trivial;
+ destruct v2; simpl; trivial;
+ destruct v3; simpl; trivial.
+ - intro Heval.
+ simpl in Heval.
+ inv Heval.
+ TrivialExists.
+ destruct v0; simpl; trivial;
+ destruct v1; simpl; trivial;
+ destruct v2; simpl; trivial.
+Qed.
+
+Lemma eval_fmaf:
+ forall al a vl v le,
+ gen_fmaf al = Some a ->
+ eval_exprlist ge sp e m le al vl ->
+ platform_builtin_sem BI_fmaf vl = Some v ->
+ exists v', eval_expr ge sp e m le a v' /\ Val.lessdef v v'.
+Proof.
+ unfold gen_fmaf.
+ intros until le.
+ intro Heval.
+ destruct (gen_fmaf_match _) in *; try discriminate.
+ all: inversion Heval; subst a; clear Heval; intro; InvEval.
+ - subst v1.
+ TrivialExists.
+ destruct v0; simpl; trivial;
+ destruct v2; simpl; trivial;
+ destruct v3; simpl; trivial.
+ - intro Heval.
+ simpl in Heval.
+ inv Heval.
+ TrivialExists.
+ destruct v0; simpl; trivial;
+ destruct v1; simpl; trivial;
+ destruct v2; simpl; trivial.
+Qed.
+
+Theorem eval_platform_builtin:
+ forall bf al a vl v le,
+ platform_builtin bf al = Some a ->
+ eval_exprlist ge sp e m le al vl ->
+ platform_builtin_sem bf vl = Some v ->
+ exists v', eval_expr ge sp e m le a v' /\ Val.lessdef v v'.
+Proof.
+ destruct bf; intros until le; intro Heval.
+ all: try (inversion Heval; subst a; clear Heval;
+ exists v; split; trivial;
+ repeat (try econstructor; try eassumption)).
+ - apply eval_fma; assumption.
+ - apply eval_fmaf; assumption.
+Qed.
+
End CMCONSTR.
diff --git a/mppa_k1c/TargetPrinter.ml b/mppa_k1c/TargetPrinter.ml
index c9822e13..5618875f 100644
--- a/mppa_k1c/TargetPrinter.ml
+++ b/mppa_k1c/TargetPrinter.ml
@@ -140,6 +140,8 @@ module Target (*: TARGET*) =
| RA -> output_string oc "$ra"
| _ -> assert false
+ let preg_asm oc ty = preg oc
+
let preg_annot = let open Asmvliw in function
| IR r -> int_reg_name r
| RA -> "$ra"
@@ -324,7 +326,7 @@ module Target (*: TARGET*) =
(P.to_int kind) (extern_atom txt) args
| EF_inline_asm(txt, sg, clob) ->
fprintf oc "%s begin inline assembly\n\t" comment;
- print_inline_asm preg oc (camlstring_of_coqstring txt) sg args res;
+ print_inline_asm preg_asm oc (camlstring_of_coqstring txt) sg args res;
fprintf oc "%s end inline assembly\n" comment
| _ ->
assert false
@@ -572,6 +574,10 @@ module Target (*: TARGET*) =
fprintf oc " maddw %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
| Pmsubw (rd, rs1, rs2) ->
fprintf oc " msbfw %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
+ | Pfmaddfw (rd, rs1, rs2) ->
+ fprintf oc " ffmaw %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
+ | Pfmsubfw (rd, rs1, rs2) ->
+ fprintf oc " ffmsw %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
| Paddl (rd, rs1, rs2) ->
fprintf oc " addd %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
@@ -613,6 +619,10 @@ module Target (*: TARGET*) =
fprintf oc " maddd %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
| Pmsubl (rd, rs1, rs2) ->
fprintf oc " msbfd %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
+ | Pfmaddfl (rd, rs1, rs2) ->
+ fprintf oc " ffmad %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
+ | Pfmsubfl (rd, rs1, rs2) ->
+ fprintf oc " ffmsd %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
| Pfaddd (rd, rs1, rs2) ->
fprintf oc " faddd %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
@@ -626,6 +636,16 @@ module Target (*: TARGET*) =
fprintf oc " fmuld %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
| Pfmulw (rd, rs1, rs2) ->
fprintf oc " fmulw %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
+ | Pfmind (rd, rs1, rs2) ->
+ fprintf oc " fmind %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
+ | Pfminw (rd, rs1, rs2) ->
+ fprintf oc " fminw %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
+ | Pfmaxd (rd, rs1, rs2) ->
+ fprintf oc " fmaxd %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
+ | Pfmaxw (rd, rs1, rs2) ->
+ fprintf oc " fmaxw %a = %a, %a\n" ireg rd ireg rs1 ireg rs2
+ | Pfinvw (rd, rs1) ->
+ fprintf oc " finvw %a = %a\n" ireg rd ireg rs1
(* Arith RRI32 instructions *)
| Pcompiw (it, rd, rs, imm) ->
diff --git a/mppa_k1c/ValueAOp.v b/mppa_k1c/ValueAOp.v
index 439138da..2c9bdf3e 100644
--- a/mppa_k1c/ValueAOp.v
+++ b/mppa_k1c/ValueAOp.v
@@ -12,7 +12,37 @@
Require Import Coqlib Compopts.
Require Import AST Integers Floats Values Memory Globalenvs.
-Require Import Op ExtValues RTL ValueDomain.
+Require Import Op ExtValues ExtFloats RTL ValueDomain.
+
+Definition minf := binop_float ExtFloat.min.
+Definition maxf := binop_float ExtFloat.max.
+Definition minfs := binop_single ExtFloat32.min.
+Definition maxfs := binop_single ExtFloat32.max.
+
+Definition ntop3 (x y z: aval) : aval := Ifptr (plub (provenance x) (plub (provenance y) (provenance z))).
+
+Definition triple_op_float (sem: float -> float -> float -> float) (x y z: aval) :=
+ match x, y, z with
+ | F a, F b, F c => F (sem a b c)
+ | _, _, _ => ntop3 x y z
+ end.
+
+Definition triple_op_single (sem: float32 -> float32 -> float32 -> float32) (x y z: aval) :=
+ match x, y, z with
+ | FS a, FS b, FS c => FS (sem a b c)
+ | _, _, _ => ntop3 x y z
+ end.
+
+Definition fmaddf := triple_op_float (fun x y z => Float.fma y z x).
+Definition fmsubf := triple_op_float (fun x y z => Float.fma (Float.neg y) z x).
+Definition fmaddfs := triple_op_single (fun x y z => Float32.fma y z x).
+Definition fmsubfs := triple_op_single (fun x y z => Float32.fma (Float32.neg y) z x).
+
+Definition invfs (y : aval) :=
+ match y with
+ | FS f => FS (ExtFloat32.inv f)
+ | _ => ntop1 y
+ end.
(** Value analysis for RISC V operators *)
@@ -235,12 +265,21 @@ Definition eval_static_operation (op: operation) (vl: list aval): aval :=
| Osubf, v1::v2::nil => subf v1 v2
| Omulf, v1::v2::nil => mulf v1 v2
| Odivf, v1::v2::nil => divf v1 v2
+ | Ominf, v1::v2::nil => minf v1 v2
+ | Omaxf, v1::v2::nil => maxf v1 v2
+ | Ofmaddf, v1::v2::v3::nil => fmaddf v1 v2 v3
+ | Ofmsubf, v1::v2::v3::nil => fmsubf v1 v2 v3
| Onegfs, v1::nil => negfs v1
| Oabsfs, v1::nil => absfs v1
| Oaddfs, v1::v2::nil => addfs v1 v2
| Osubfs, v1::v2::nil => subfs v1 v2
| Omulfs, v1::v2::nil => mulfs v1 v2
| Odivfs, v1::v2::nil => divfs v1 v2
+ | Ominfs, v1::v2::nil => minfs v1 v2
+ | Omaxfs, v1::v2::nil => maxfs v1 v2
+ | Oinvfs, v1::nil => invfs v1
+ | Ofmaddfs, v1::v2::v3::nil => fmaddfs v1 v2 v3
+ | Ofmsubfs, v1::v2::v3::nil => fmsubfs v1 v2 v3
| Osingleoffloat, v1::nil => singleoffloat v1
| Ofloatofsingle, v1::nil => floatofsingle v1
| Ointoffloat, v1::nil => intoffloat v1
@@ -278,6 +317,99 @@ Hypothesis GENV: genv_match bc ge.
Variable sp: block.
Hypothesis STACK: bc sp = BCstack.
+Lemma minf_sound:
+ forall v x w y, vmatch bc v x -> vmatch bc w y -> vmatch bc (ExtValues.minf v w) (minf x y).
+Proof.
+ apply (binop_float_sound bc ExtFloat.min); assumption.
+Qed.
+
+Lemma maxf_sound:
+ forall v x w y, vmatch bc v x -> vmatch bc w y -> vmatch bc (ExtValues.maxf v w) (maxf x y).
+Proof.
+ apply (binop_float_sound bc ExtFloat.max); assumption.
+Qed.
+
+Lemma minfs_sound:
+ forall v x w y, vmatch bc v x -> vmatch bc w y -> vmatch bc (ExtValues.minfs v w) (minfs x y).
+Proof.
+ apply (binop_single_sound bc ExtFloat32.min); assumption.
+Qed.
+
+Lemma maxfs_sound:
+ forall v x w y, vmatch bc v x -> vmatch bc w y -> vmatch bc (ExtValues.maxfs v w) (maxfs x y).
+Proof.
+ apply (binop_single_sound bc ExtFloat32.max); assumption.
+Qed.
+
+Lemma invfs_sound:
+ forall v x, vmatch bc v x -> vmatch bc (ExtValues.invfs v) (invfs x).
+Proof.
+ intros v x;
+ intro MATCH;
+ inversion MATCH;
+ simpl;
+ constructor.
+Qed.
+
+Lemma triple_op_float_sound:
+ forall f a x b y c z,
+ vmatch bc a x -> vmatch bc b y -> vmatch bc c z ->
+ vmatch bc (ExtValues.triple_op_float f a b c)
+ (triple_op_float f x y z).
+Proof.
+ intros until z.
+ intros Hax Hby Hcz.
+ inv Hax; simpl; try constructor;
+ inv Hby; simpl; try constructor;
+ inv Hcz; simpl; try constructor.
+Qed.
+
+Lemma triple_op_single_sound:
+ forall f a x b y c z,
+ vmatch bc a x -> vmatch bc b y -> vmatch bc c z ->
+ vmatch bc (ExtValues.triple_op_single f a b c)
+ (triple_op_single f x y z).
+Proof.
+ intros until z.
+ intros Hax Hby Hcz.
+ inv Hax; simpl; try constructor;
+ inv Hby; simpl; try constructor;
+ inv Hcz; simpl; try constructor.
+Qed.
+
+Lemma fmaddf_sound :
+ forall a x b y c z, vmatch bc a x -> vmatch bc b y -> vmatch bc c z ->
+ vmatch bc (ExtValues.fmaddf a b c) (fmaddf x y z).
+Proof.
+ intros. unfold ExtValues.fmaddf, fmaddf.
+ apply triple_op_float_sound; assumption.
+Qed.
+
+Lemma fmaddfs_sound :
+ forall a x b y c z, vmatch bc a x -> vmatch bc b y -> vmatch bc c z ->
+ vmatch bc (ExtValues.fmaddfs a b c) (fmaddfs x y z).
+Proof.
+ intros. unfold ExtValues.fmaddfs, fmaddfs.
+ apply triple_op_single_sound; assumption.
+Qed.
+
+Lemma fmsubf_sound :
+ forall a x b y c z, vmatch bc a x -> vmatch bc b y -> vmatch bc c z ->
+ vmatch bc (ExtValues.fmsubf a b c) (fmsubf x y z).
+Proof.
+ intros. unfold ExtValues.fmsubf, fmsubf.
+ apply triple_op_float_sound; assumption.
+Qed.
+
+Lemma fmsubfs_sound :
+ forall a x b y c z, vmatch bc a x -> vmatch bc b y -> vmatch bc c z ->
+ vmatch bc (ExtValues.fmsubfs a b c) (fmsubfs x y z).
+Proof.
+ intros. unfold ExtValues.fmsubfs, fmsubfs.
+ apply triple_op_single_sound; assumption.
+Qed.
+Hint Resolve minf_sound maxf_sound minfs_sound maxfs_sound invfs_sound fmaddf_sound fmaddfs_sound fmsubf_sound fmsubfs_sound : va.
+
Theorem eval_static_condition_sound:
forall cond vargs m aargs,
list_forall2 (vmatch bc) vargs aargs ->
diff --git a/mppa_k1c/lib/Machblock.v b/mppa_k1c/lib/Machblock.v
index 30393fd5..2759c49d 100644
--- a/mppa_k1c/lib/Machblock.v
+++ b/mppa_k1c/lib/Machblock.v
@@ -14,7 +14,7 @@ Require Stacklayout.
Require Import Mach.
Require Import Linking.
-(** instructions "basiques" (ie non control-flow) *)
+(** basic instructions (ie no control-flow) *)
Inductive basic_inst: Type :=
| MBgetstack: ptrofs -> typ -> mreg -> basic_inst
| MBsetstack: mreg -> ptrofs -> typ -> basic_inst
@@ -26,7 +26,7 @@ Inductive basic_inst: Type :=
Definition bblock_body := list basic_inst.
-(** instructions de control flow *)
+(** control flow instructions *)
Inductive control_flow_inst: Type :=
| MBcall: signature -> mreg + ident -> control_flow_inst
| MBtailcall: signature -> mreg + ident -> control_flow_inst
diff --git a/mppa_k1c/lib/Machblockgen.v b/mppa_k1c/lib/Machblockgen.v
index 4dfc309e..db48934e 100644
--- a/mppa_k1c/lib/Machblockgen.v
+++ b/mppa_k1c/lib/Machblockgen.v
@@ -57,12 +57,9 @@ Definition add_to_new_bblock (i:Machblock_inst) : bblock :=
| MB_cfi i => cfi_bblock i
end.
-(* ajout d'une instruction en début d'une liste de blocks *)
-(* Soit /1\ ajout en tête de block, soit /2\ ajout dans un nouveau block*)
-(* bl est vide -> /2\ *)
-(* cfi -> /2\ (ajout dans exit)*)
-(* basic -> /1\ si header vide, /2\ si a un header *)
-(* label -> /1\ (dans header)*)
+(** Adding an instruction to the beginning of a bblock list
+ * Either adding the instruction to the head of the list,
+ * or create a new bblock with the instruction *)
Definition add_to_code (i:Machblock_inst) (bl:code) : code :=
match bl with
| bh::bl0 => match i with
@@ -86,8 +83,6 @@ Fixpoint trans_code_rev (c: Mach.code) (bl:code) : code :=
Function trans_code (c: Mach.code) : code :=
trans_code_rev (List.rev_append c nil) nil.
-
-(* à finir pour passer des Mach.function au function, etc. *)
Definition transf_function (f: Mach.function) : function :=
{| fn_sig:=Mach.fn_sig f;
fn_code:=trans_code (Mach.fn_code f);
@@ -103,7 +98,7 @@ Definition transf_program (src: Mach.program) : program :=
transform_program transf_fundef src.
-(** Abstraction de trans_code *)
+(** Abstracting trans_code *)
Inductive is_end_block: Machblock_inst -> code -> Prop :=
| End_empty mbi: is_end_block mbi nil
diff --git a/pg b/pg
index 28926baa..398d618f 100755
--- a/pg
+++ b/pg
@@ -1,10 +1,7 @@
#!/bin/sh
-# Start Proof General with the right -I options
+# Start Proof General with the right Coq version
# Use the Makefile to rebuild dependencies if needed
-# Recompile the modified file after coqide editing
-
-PWD=`pwd`
-INCLUDES=`make print-includes`
+# Recompile the modified file after editing
make -q ${1}o || {
make -n ${1}o | grep -v "\\b${1}\\b" | \
@@ -15,16 +12,5 @@ make -q ${1}o || {
COQPROGNAME="${COQBIN}coqtop"
-COQPROGARGS=""
-for arg in $INCLUDES; do
- case "$arg" in
- -I|-R|-as|compcert*)
- COQPROGARGS="$COQPROGARGS \"$arg\"";;
- *)
- COQPROGARGS="$COQPROGARGS \"$PWD/$arg\"";;
- esac
-done
-
-emacs --eval "(setq coq-prog-name \"$COQPROGNAME\")" \
- --eval "(setq coq-prog-args '($COQPROGARGS))" $1 \
+emacs --eval "(setq coq-prog-name \"$COQPROGNAME\")" $1 \
&& make ${1}o
diff --git a/powerpc/Archi.v b/powerpc/Archi.v
index d792e4fe..ab348c14 100644
--- a/powerpc/Archi.v
+++ b/powerpc/Archi.v
@@ -16,7 +16,7 @@
(** Architecture-dependent parameters for PowerPC *)
-Require Import ZArith.
+Require Import ZArith List.
(*From Flocq*)
Require Import Binary Bits.
@@ -37,24 +37,33 @@ Proof.
reflexivity.
Qed.
-Definition default_nan_64 : { x : binary64 | is_nan _ _ x = true } :=
- exist _ (B754_nan 53 1024 false (iter_nat 51 _ xO xH) (eq_refl true)) (eq_refl true).
+Definition default_nan_64 := (false, iter_nat 51 _ xO xH).
+Definition default_nan_32 := (false, iter_nat 22 _ xO xH).
-Definition choose_binop_pl_64 (pl1 pl2 : positive) :=
- false. (**r always choose first NaN *)
+(* Always choose the first NaN argument, if any *)
-Definition default_nan_32 : { x : binary32 | is_nan _ _ x = true } :=
- exist _ (B754_nan 24 128 false (iter_nat 22 _ xO xH) (eq_refl true)) (eq_refl true).
+Definition choose_nan_64 (l: list (bool * positive)) : bool * positive :=
+ match l with nil => default_nan_64 | n :: _ => n end.
-Definition choose_binop_pl_32 (pl1 pl2 : positive) :=
- false. (**r always choose first NaN *)
+Definition choose_nan_32 (l: list (bool * positive)) : bool * positive :=
+ match l with nil => default_nan_32 | n :: _ => n end.
-Definition fpu_returns_default_qNaN := false.
+Lemma choose_nan_64_idem: forall n,
+ choose_nan_64 (n :: n :: nil) = choose_nan_64 (n :: nil).
+Proof. auto. Qed.
+
+Lemma choose_nan_32_idem: forall n,
+ choose_nan_32 (n :: n :: nil) = choose_nan_32 (n :: nil).
+Proof. auto. Qed.
+
+Definition fma_order {A: Type} (x y z: A) := (x, z, y).
+
+Definition fma_invalid_mul_is_nan := false.
Definition float_of_single_preserves_sNaN := true.
Global Opaque ptr64 big_endian splitlong
- default_nan_64 choose_binop_pl_64
- default_nan_32 choose_binop_pl_32
- fpu_returns_default_qNaN
+ default_nan_64 choose_nan_64
+ default_nan_32 choose_nan_32
+ fma_order fma_invalid_mul_is_nan
float_of_single_preserves_sNaN.
diff --git a/powerpc/Asmexpand.ml b/powerpc/Asmexpand.ml
index 415b6651..5ca4c611 100644
--- a/powerpc/Asmexpand.ml
+++ b/powerpc/Asmexpand.ml
@@ -554,6 +554,26 @@ let expand_builtin_inline name args res =
emit (Plabel lbl2)
| "__builtin_cmpb", [BA(IR a1); BA(IR a2)], BR(IR res) ->
emit (Pcmpb (res,a1,a2))
+ | "__builtin_bswap64", [BA_splitlong(BA(IR ah), BA(IR al))],
+ BR_splitlong(BR(IR rh), BR(IR rl))->
+ assert (not Archi.ppc64);
+ emit (Pstwu(ah, Cint _m8, GPR1));
+ emit (Pcfi_adjust _8);
+ emit (Pstwu(al, Cint _m8, GPR1));
+ emit (Pcfi_adjust _8);
+ emit (Plwbrx(rh, GPR0, GPR1));
+ emit (Paddi(GPR1, GPR1, Cint _8));
+ emit (Pcfi_adjust _m8);
+ emit (Plwbrx(rl, GPR0, GPR1));
+ emit (Paddi(GPR1, GPR1, Cint _8));
+ emit (Pcfi_adjust _m8)
+ | "__builtin_bswap64", [BA(IR a1)], BR(IR res) ->
+ assert (Archi.ppc64);
+ emit (Pstdu(a1, Cint _m8, GPR1));
+ emit (Pcfi_adjust _8);
+ emit (Pldbrx(res, GPR0, GPR1));
+ emit (Paddi(GPR1, GPR1, Cint _8));
+ emit (Pcfi_adjust _m8)
| ("__builtin_bswap" | "__builtin_bswap32"), [BA(IR a1)], BR(IR res) ->
emit (Pstwu(a1, Cint _m8, GPR1));
emit (Pcfi_adjust _8);
diff --git a/powerpc/Builtins1.v b/powerpc/Builtins1.v
new file mode 100644
index 00000000..f6e643d2
--- /dev/null
+++ b/powerpc/Builtins1.v
@@ -0,0 +1,33 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, Collège de France and Inria Paris *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the GNU General Public License as published by *)
+(* the Free Software Foundation, either version 2 of the License, or *)
+(* (at your option) any later version. This file is also distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Platform-specific built-in functions *)
+
+Require Import String Coqlib.
+Require Import AST Integers Floats Values.
+Require Import Builtins0.
+
+Inductive platform_builtin : Type := .
+
+Local Open Scope string_scope.
+
+Definition platform_builtin_table : list (string * platform_builtin) :=
+ nil.
+
+Definition platform_builtin_sig (b: platform_builtin) : signature :=
+ match b with end.
+
+Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (proj_sig_res (platform_builtin_sig b)) :=
+ match b with end.
diff --git a/powerpc/CBuiltins.ml b/powerpc/CBuiltins.ml
index 11b7aef9..e29a41f1 100644
--- a/powerpc/CBuiltins.ml
+++ b/powerpc/CBuiltins.ml
@@ -18,11 +18,11 @@
open C
let builtins = {
- Builtins.typedefs = [
+ builtin_typedefs = [
"__builtin_va_list",
TArray(TInt(IUInt, []), Some 3L, [])
];
- Builtins.functions = [
+ builtin_functions = [
(* Integer arithmetic *)
"__builtin_mulhw",
(TInt(IInt, []), [TInt(IInt, []); TInt(IInt, [])], false);
diff --git a/powerpc/ConstpropOp.vp b/powerpc/ConstpropOp.vp
index 2d492b66..8e90a849 100644
--- a/powerpc/ConstpropOp.vp
+++ b/powerpc/ConstpropOp.vp
@@ -16,13 +16,14 @@
Require Import Coqlib Compopts.
Require Import AST Integers Floats.
Require Import Op Registers.
-Require Import ValueDomain.
+Require Import ValueDomain ValueAOp.
(** * Converting known values to constants *)
Definition const_for_result (a: aval) : option operation :=
match a with
| I n => Some(Ointconst n)
+ | L n => if Archi.ppc64 then Some (Olongconst n) else None
| F n => if Compopts.generate_float_constants tt then Some(Ofloatconst n) else None
| FS n => if Compopts.generate_float_constants tt then Some(Osingleconst n) else None
| Ptr(Gl id ofs) => Some (Oaddrsymbol id ofs)
@@ -95,6 +96,15 @@ Nondetfunction make_cmp (c: condition) (args: list reg) (vl: list aval) :=
make_cmp_base c args vl
end.
+Definition make_select (c: condition) (ty: typ)
+ (r1 r2: reg) (args: list reg) (vl: list aval) :=
+ match resolve_branch (eval_static_condition c vl) with
+ | Some b => (Omove, (if b then r1 else r2) :: nil)
+ | None =>
+ let (c', args') := cond_strength_reduction c args vl in
+ (Osel c' ty, r1 :: r2 :: args')
+ end.
+
Definition make_addimm (n: int) (r: reg) :=
if Int.eq n Int.zero
then (Omove, r :: nil)
@@ -303,6 +313,7 @@ Nondetfunction op_strength_reduction
| Oshrl, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrlimm n2 r1 r2
| Oshrlu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrluimm n2 r1 r2
| Ocmp c, args, vl => make_cmp c args vl
+ | Osel c ty, r1 :: r2 :: args, v1 :: v2 :: vl => make_select c ty r1 r2 args vl
| Omulf, r1 :: r2 :: nil, v1 :: F n2 :: nil => make_mulfimm n2 r1 r1 r2
| Omulf, r1 :: r2 :: nil, F n1 :: v2 :: nil => make_mulfimm n1 r2 r1 r2
| Omulfs, r1 :: r2 :: nil, v1 :: FS n2 :: nil => make_mulfsimm n2 r1 r1 r2
diff --git a/powerpc/ConstpropOpproof.v b/powerpc/ConstpropOpproof.v
index fe061e5b..8687b056 100644
--- a/powerpc/ConstpropOpproof.v
+++ b/powerpc/ConstpropOpproof.v
@@ -14,7 +14,7 @@
Require Import Coqlib Compopts.
Require Import Integers Floats Values Memory Globalenvs Events.
-Require Import Op Registers RTL ValueDomain.
+Require Import Op Registers RTL ValueDomain ValueAOp ValueAnalysis.
Require Import ConstpropOp.
Local Transparent Archi.ptr64.
@@ -101,6 +101,8 @@ Proof.
destruct a; inv H; SimplVM.
- (* integer *)
exists (Vint n); auto.
+- (* long *)
+ destruct (Archi.ppc64); inv H2. exists (Vlong n); auto.
- (* float *)
destruct (generate_float_constants tt); inv H2. exists (Vfloat f); auto.
- (* single *)
@@ -211,6 +213,28 @@ Proof.
- apply make_cmp_base_correct; auto.
Qed.
+Lemma make_select_correct:
+ forall c ty r1 r2 args vl,
+ vl = map (fun r => AE.get r ae) args ->
+ let (op', args') := make_select c ty r1 r2 args vl in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op' rs##args' m = Some v
+ /\ Val.lessdef (Val.select (eval_condition c rs##args m) rs#r1 rs#r2 ty) v.
+Proof.
+ unfold make_select; intros.
+ destruct (resolve_branch (eval_static_condition c vl)) as [b|] eqn:RB.
+- exists (if b then rs#r1 else rs#r2); split.
++ simpl. destruct b; auto.
++ destruct (eval_condition c rs##args m) as [b'|] eqn:EC; simpl; auto.
+ assert (b = b').
+ { eapply resolve_branch_sound; eauto.
+ rewrite <- EC. apply eval_static_condition_sound with bc.
+ subst vl. exact (aregs_sound _ _ _ args MATCH). }
+ subst b'. apply Val.lessdef_normalize.
+- generalize (cond_strength_reduction_correct c args vl H).
+ destruct (cond_strength_reduction c args vl) as [cond' args']; intros EQ.
+ econstructor; split. simpl; eauto. rewrite EQ; auto.
+Qed.
+
Lemma make_addimm_correct:
forall n r,
let (op, args) := make_addimm n r in
@@ -715,6 +739,8 @@ Proof.
InvApproxRegs; SimplVM; inv H0. apply make_shrluimm_correct; auto.
(* cmp *)
inv H0. apply make_cmp_correct; auto.
+(* select *)
+ inv H0. apply make_select_correct; congruence.
(* mulf *)
InvApproxRegs; SimplVM; inv H0. rewrite <- H2. apply make_mulfimm_correct; auto.
InvApproxRegs; SimplVM; inv H0. fold (Val.mulf (Vfloat n1) rs#r2).
diff --git a/powerpc/SelectOp.vp b/powerpc/SelectOp.vp
index b1cac124..50b1bdd6 100644
--- a/powerpc/SelectOp.vp
+++ b/powerpc/SelectOp.vp
@@ -38,12 +38,8 @@
Require Import Coqlib.
Require Import Compopts.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Op.
-Require Import CminorSel.
-Require Import OpHelpers.
+Require Import AST Integers Floats Builtins.
+Require Import Op OpHelpers CminorSel.
Require Archi.
Local Open Scope cminorsel_scope.
@@ -574,3 +570,8 @@ Definition divf_base (e1: expr) (e2: expr) :=
Definition divfs_base (e1: expr) (e2: expr) :=
Eop Odivfs (e1 ::: e2 ::: Enil).
+
+(** Platform-specific known builtins *)
+
+Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr :=
+ None.
diff --git a/powerpc/SelectOpproof.v b/powerpc/SelectOpproof.v
index 92852d36..b0cd70a4 100644
--- a/powerpc/SelectOpproof.v
+++ b/powerpc/SelectOpproof.v
@@ -13,17 +13,10 @@
(** Correctness of instruction selection for operators *)
Require Import Coqlib.
-Require Import Maps.
+Require Import AST Integers Floats.
+Require Import Values Memory Builtins Globalenvs.
+Require Import Cminor Op CminorSel.
Require Import Compopts.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Cminor.
-Require Import Op.
-Require Import CminorSel.
Require Import SelectOp.
Require Import OpHelpers.
Require Import OpHelpersproof.
@@ -1089,4 +1082,17 @@ Proof.
intros; unfold divfs_base.
TrivialExists.
Qed.
+
+(** Platform-specific known builtins *)
+
+Theorem eval_platform_builtin:
+ forall bf al a vl v le,
+ platform_builtin bf al = Some a ->
+ eval_exprlist ge sp e m le al vl ->
+ platform_builtin_sem bf vl = Some v ->
+ exists v', eval_expr ge sp e m le a v' /\ Val.lessdef v v'.
+Proof.
+ intros. discriminate.
+Qed.
+
End CMCONSTR.
diff --git a/powerpc/TargetPrinter.ml b/powerpc/TargetPrinter.ml
index a1338561..0f608d25 100644
--- a/powerpc/TargetPrinter.ml
+++ b/powerpc/TargetPrinter.ml
@@ -340,7 +340,7 @@ module Target (System : SYSTEM):TARGET =
let ireg_or_zero oc r =
if r = GPR0 then output_string oc "0" else ireg oc r
- let preg oc = function
+ let preg_asm oc ty = function
| IR r -> ireg oc r
| FR r -> freg oc r
| _ -> assert false
@@ -863,7 +863,7 @@ module Target (System : SYSTEM):TARGET =
(P.to_int kind) (extern_atom txt) args
| EF_inline_asm(txt, sg, clob) ->
fprintf oc "%s begin inline assembly\n\t" comment;
- print_inline_asm preg oc (camlstring_of_coqstring txt) sg args res;
+ print_inline_asm preg_asm oc (camlstring_of_coqstring txt) sg args res;
fprintf oc "%s end inline assembly\n" comment
| _ ->
assert false
diff --git a/riscV/Archi.v b/riscV/Archi.v
index 3758d686..61d129d0 100644
--- a/riscV/Archi.v
+++ b/riscV/Archi.v
@@ -16,7 +16,7 @@
(** Architecture-dependent parameters for RISC-V *)
-Require Import ZArith.
+Require Import ZArith List.
(*From Flocq*)
Require Import Binary Bits.
@@ -40,26 +40,33 @@ Qed.
except the MSB, a.k.a. the quiet bit."
Exceptions are operations manipulating signs. *)
-Definition default_nan_64 : { x : binary64 | is_nan _ _ x = true } :=
- exist _ (B754_nan 53 1024 false (iter_nat 51 _ xO xH) (eq_refl true)) (eq_refl true).
+Definition default_nan_64 := (false, iter_nat 51 _ xO xH).
+Definition default_nan_32 := (false, iter_nat 22 _ xO xH).
-Definition choose_binop_pl_64 (pl1 pl2 : positive) :=
- false. (**r irrelevant *)
+Definition choose_nan_64 (l: list (bool * positive)) : bool * positive :=
+ default_nan_64.
-Definition default_nan_32 : { x : binary32 | is_nan _ _ x = true } :=
- exist _ (B754_nan 24 128 false (iter_nat 22 _ xO xH) (eq_refl true)) (eq_refl true).
+Definition choose_nan_32 (l: list (bool * positive)) : bool * positive :=
+ default_nan_32.
-Definition choose_binop_pl_32 (pl1 pl2 : positive) :=
- false. (**r irrelevant *)
+Lemma choose_nan_64_idem: forall n,
+ choose_nan_64 (n :: n :: nil) = choose_nan_64 (n :: nil).
+Proof. auto. Qed.
-Definition fpu_returns_default_qNaN := true.
+Lemma choose_nan_32_idem: forall n,
+ choose_nan_32 (n :: n :: nil) = choose_nan_32 (n :: nil).
+Proof. auto. Qed.
+
+Definition fma_order {A: Type} (x y z: A) := (x, y, z).
+
+Definition fma_invalid_mul_is_nan := false.
Definition float_of_single_preserves_sNaN := false.
Global Opaque ptr64 big_endian splitlong
- default_nan_64 choose_binop_pl_64
- default_nan_32 choose_binop_pl_32
- fpu_returns_default_qNaN
+ default_nan_64 choose_nan_64
+ default_nan_32 choose_nan_32
+ fma_order fma_invalid_mul_is_nan
float_of_single_preserves_sNaN.
(** Whether to generate position-independent code or not *)
diff --git a/riscV/Builtins1.v b/riscV/Builtins1.v
new file mode 100644
index 00000000..f6e643d2
--- /dev/null
+++ b/riscV/Builtins1.v
@@ -0,0 +1,33 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, Collège de France and Inria Paris *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the GNU General Public License as published by *)
+(* the Free Software Foundation, either version 2 of the License, or *)
+(* (at your option) any later version. This file is also distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Platform-specific built-in functions *)
+
+Require Import String Coqlib.
+Require Import AST Integers Floats Values.
+Require Import Builtins0.
+
+Inductive platform_builtin : Type := .
+
+Local Open Scope string_scope.
+
+Definition platform_builtin_table : list (string * platform_builtin) :=
+ nil.
+
+Definition platform_builtin_sig (b: platform_builtin) : signature :=
+ match b with end.
+
+Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (proj_sig_res (platform_builtin_sig b)) :=
+ match b with end.
diff --git a/riscV/CBuiltins.ml b/riscV/CBuiltins.ml
index 0c981d11..a2087cb7 100644
--- a/riscV/CBuiltins.ml
+++ b/riscV/CBuiltins.ml
@@ -18,16 +18,13 @@
open C
let builtins = {
- Builtins.typedefs = [
+ builtin_typedefs = [
"__builtin_va_list", TPtr(TVoid [], [])
];
- Builtins.functions = [
+ builtin_functions = [
(* Synchronization *)
"__builtin_fence",
(TVoid [], [], false);
- (* Integer arithmetic *)
- "__builtin_bswap64",
- (TInt(IULongLong, []), [TInt(IULongLong, [])], false);
(* Float arithmetic *)
"__builtin_fmadd",
(TFloat(FDouble, []),
diff --git a/riscV/SelectOp.vp b/riscV/SelectOp.vp
index ba7c5664..e9920e46 100644
--- a/riscV/SelectOp.vp
+++ b/riscV/SelectOp.vp
@@ -44,11 +44,8 @@
Require Archi.
Require Import Coqlib.
Require Import Compopts.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Op.
-Require Import CminorSel.
+Require Import AST Integers Floats Builtins.
+Require Import Op CminorSel.
Local Open Scope cminorsel_scope.
@@ -461,3 +458,8 @@ Definition divf_base (e1: expr) (e2: expr) :=
Definition divfs_base (e1: expr) (e2: expr) :=
Eop Odivfs (e1 ::: e2 ::: Enil).
+
+(** Platform-specific known builtins *)
+
+Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr :=
+ None.
diff --git a/riscV/SelectOpproof.v b/riscV/SelectOpproof.v
index 46cc1bbc..52a535fb 100644
--- a/riscV/SelectOpproof.v
+++ b/riscV/SelectOpproof.v
@@ -17,18 +17,10 @@
(** Correctness of instruction selection for operators *)
-Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import Zbits.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Cminor.
-Require Import Op.
-Require Import CminorSel.
+Require Import Coqlib Zbits.
+Require Import AST Integers Floats.
+Require Import Values Memory Builtins Globalenvs.
+Require Import Cminor Op CminorSel.
Require Import SelectOp.
Require Import OpHelpers.
Require Import OpHelpersproof.
@@ -941,7 +933,6 @@ Proof.
- constructor; auto.
Qed.
-
(* floating-point division without HELPERS *)
Theorem eval_divf_base:
forall le a b x y,
@@ -962,4 +953,17 @@ Proof.
intros; unfold divfs_base.
TrivialExists.
Qed.
+
+(** Platform-specific known builtins *)
+
+Theorem eval_platform_builtin:
+ forall bf al a vl v le,
+ platform_builtin bf al = Some a ->
+ eval_exprlist ge sp e m le al vl ->
+ platform_builtin_sem bf vl = Some v ->
+ exists v', eval_expr ge sp e m le a v' /\ Val.lessdef v v'.
+Proof.
+ intros. discriminate.
+Qed.
+
End CMCONSTR.
diff --git a/riscV/TargetPrinter.ml b/riscV/TargetPrinter.ml
index 92df7a76..64bcea4c 100644
--- a/riscV/TargetPrinter.ml
+++ b/riscV/TargetPrinter.ml
@@ -93,7 +93,7 @@ module Target : TARGET =
| X0 -> output_string oc "x0"
| X r -> ireg oc r
- let preg oc = function
+ let preg_asm oc ty = function
| IR r -> ireg oc r
| FR r -> freg oc r
| _ -> assert false
@@ -582,7 +582,7 @@ module Target : TARGET =
(P.to_int kind) (extern_atom txt) args
| EF_inline_asm(txt, sg, clob) ->
fprintf oc "%s begin inline assembly\n\t" comment;
- print_inline_asm preg oc (camlstring_of_coqstring txt) sg args res;
+ print_inline_asm preg_asm oc (camlstring_of_coqstring txt) sg args res;
fprintf oc "%s end inline assembly\n" comment
| _ ->
assert false
diff --git a/runtime/include/math.h b/runtime/include/math.h
index 805cc8e7..060968c8 100644
--- a/runtime/include/math.h
+++ b/runtime/include/math.h
@@ -3,5 +3,16 @@
#define isfinite(__y) (fpclassify((__y)) >= FP_ZERO)
+#ifndef COMPCERT_NO_FP_MACROS
+#define fmin(x, y) __builtin_fmin((x),(y))
+#define fmax(x, y) __builtin_fmax((x),(y))
+#define fminf(x, y) __builtin_fminf((x),(y))
+#define fmaxf(x, y) __builtin_fmaxf((x),(y))
+#define fabs(x) __builtin_fabs((x))
+#define fabsf(x) __builtin_fabsf((x))
+#define fma(x, y, z) __builtin_fma((x),(y),(z))
+#define fmaf(x, y, z) __builtin_fmaf((x),(y),(z))
+#endif
+
#include_next <math.h>
#endif
diff --git a/runtime/powerpc/i64_dtos.s b/runtime/powerpc/i64_dtos.s
deleted file mode 100644
index 85c60b27..00000000
--- a/runtime/powerpc/i64_dtos.s
+++ /dev/null
@@ -1,100 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-### Conversion from double float to signed long
-
- .balign 16
- .globl __compcert_i64_dtos
-__compcert_i64_dtos:
- stfdu f1, -16(r1) # extract LO (r4) and HI (r3) halves of double
- lwz r3, 0(r1)
- lwz r4, 4(r1)
- addi r1, r1, 16
- srawi r10, r3, 31 # save sign of double in r10
- # extract unbiased exponent ((HI & 0x7FF00000) >> 20) - (1023 + 52)
- rlwinm r5, r3, 12, 21, 31
- addi r5, r5, -1075
- # check range of exponent
- cmpwi r5, -52 # if EXP < -52, abs(double) is < 1.0
- blt 1f
- cmpwi r5, 11 # if EXP >= 63 - 52, abs(double) is >= 2^63
- bge 2f
- # extract true mantissa
- rlwinm r3, r3, 0, 12, 31 # HI &= ~0xFFF00000
- oris r3, r3, 0x10 # HI |= 0x00100000
- # shift it appropriately
- cmpwi r5, 0
- blt 3f
- # if EXP >= 0, shift left by EXP. Note that EXP < 11.
- subfic r6, r5, 32 # r6 = 32 - EXP
- slw r3, r3, r5
- srw r0, r4, r6
- or r3, r3, r0
- slw r4, r4, r5
- b 4f
- # if EXP < 0, shift right by -EXP. Note that -EXP <= 52 but can be >= 32.
-3: subfic r5, r5, 0 # r5 = -EXP = shift amount
- subfic r6, r5, 32 # r6 = 32 - amount
- addi r7, r5, -32 # r7 = amount - 32 (see i64_shr.s)
- srw r4, r4, r5
- slw r0, r3, r6
- or r4, r4, r0
- srw r0, r3, r7
- or r4, r4, r0
- srw r3, r3, r5
- # apply sign to result
-4: xor r4, r4, r10
- xor r3, r3, r10
- subfc r4, r10, r4
- subfe r3, r10, r3
- blr
- # Special cases
-1: li r3, 0 # result is 0
- li r4, 0
- blr
-2: li r4, -1 # result is MAX_SINT or MIN_SINT
- bge 5f # depending on sign
- li r4, -1 # result is MAX_SINT = 0x7FFF_FFFF
- srwi r3, r4, 1
- blr
-5: lis r3, 0x8000 # result is MIN_SINT = 0x8000_0000
- li r4, 0
- blr
- .type __compcert_i64_dtos, @function
- .size __compcert_i64_dtos, .-__compcert_i64_dtos
- \ No newline at end of file
diff --git a/runtime/powerpc/i64_dtou.s b/runtime/powerpc/i64_dtou.s
deleted file mode 100644
index 67a721d4..00000000
--- a/runtime/powerpc/i64_dtou.s
+++ /dev/null
@@ -1,92 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-### Conversion from double float to unsigned long
-
- .balign 16
- .globl __compcert_i64_dtou
-__compcert_i64_dtou:
- stfdu f1, -16(r1) # extract LO (r4) and HI (r3) halves of double
- lwz r3, 0(r1)
- lwz r4, 4(r1)
- addi r1, r1, 16
- cmpwi r3, 0 # is double < 0?
- blt 1f # then it converts to 0
- # extract unbiased exponent ((HI & 0x7FF00000) >> 20) - (1023 + 52)
- rlwinm r5, r3, 12, 21, 31
- addi r5, r5, -1075
- # check range of exponent
- cmpwi r5, -52 # if EXP < -52, double is < 1.0
- blt 1f
- cmpwi r5, 12 # if EXP >= 64 - 52, double is >= 2^64
- bge 2f
- # extract true mantissa
- rlwinm r3, r3, 0, 12, 31 # HI &= ~0xFFF00000
- oris r3, r3, 0x10 # HI |= 0x00100000
- # shift it appropriately
- cmpwi r5, 0
- blt 3f
- # if EXP >= 0, shift left by EXP. Note that EXP < 12.
- subfic r6, r5, 32 # r6 = 32 - EXP
- slw r3, r3, r5
- srw r0, r4, r6
- or r3, r3, r0
- slw r4, r4, r5
- blr
- # if EXP < 0, shift right by -EXP. Note that -EXP <= 52 but can be >= 32.
-3: subfic r5, r5, 0 # r5 = -EXP = shift amount
- subfic r6, r5, 32 # r6 = 32 - amount
- addi r7, r5, -32 # r7 = amount - 32 (see i64_shr.s)
- srw r4, r4, r5
- slw r0, r3, r6
- or r4, r4, r0
- srw r0, r3, r7
- or r4, r4, r0
- srw r3, r3, r5
- blr
- # Special cases
-1: li r3, 0 # result is 0
- li r4, 0
- blr
-2: li r3, -1 # result is MAX_UINT
- li r4, -1
- blr
- .type __compcert_i64_dtou, @function
- .size __compcert_i64_dtou, .-__compcert_i64_dtou
-
- \ No newline at end of file
diff --git a/runtime/powerpc/i64_sar.s b/runtime/powerpc/i64_sar.s
deleted file mode 100644
index c7da448f..00000000
--- a/runtime/powerpc/i64_sar.s
+++ /dev/null
@@ -1,60 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-# Shift right signed
-
- .balign 16
- .globl __compcert_i64_sar
-__compcert_i64_sar:
- andi. r5, r5, 63 # take amount modulo 64
- cmpwi r5, 32
- bge 1f # amount < 32?
- subfic r6, r5, 32 # r6 = 32 - amount
- srw r4, r4, r5 # RH = XH >>s amount
- slw r0, r3, r6 # RL = XL >>u amount | XH << (32 - amount)
- or r4, r4, r0
- sraw r3, r3, r5
- blr
-1: addi r6, r5, -32 # amount >= 32
- sraw r4, r3, r6 # RL = XH >>s (amount - 32)
- srawi r3, r3, 31 # RL = sign extension of XH
- blr
- .type __compcert_i64_sar, @function
- .size __compcert_i64_sar, .-__compcert_i64_sar
-
- \ No newline at end of file
diff --git a/runtime/powerpc/i64_sdiv.s b/runtime/powerpc/i64_sdiv.s
deleted file mode 100644
index 9787ea3b..00000000
--- a/runtime/powerpc/i64_sdiv.s
+++ /dev/null
@@ -1,71 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-### Signed division
-
- .balign 16
- .globl __compcert_i64_sdiv
-__compcert_i64_sdiv:
- mflr r0
- stw r0, 4(r1) # save return address in caller's frame
- xor r0, r3, r5 # compute sign of result (top bit)
- mtctr r0 # save it in CTR (why not?)
- srawi r0, r3, 31 # take absolute value of N
- xor r4, r4, r0 # (i.e. N = N ^ r0 - r0,
- xor r3, r3, r0 # where r0 = 0 if N >= 0 and r0 = -1 if N < 0)
- subfc r4, r0, r4
- subfe r3, r0, r3
- srawi r0, r5, 31 # take absolute value of D
- xor r6, r6, r0 # (same trick)
- xor r5, r5, r0
- subfc r6, r0, r6
- subfe r5, r0, r5
- bl __compcert_i64_udivmod # do unsigned division
- lwz r0, 4(r1)
- mtlr r0 # restore return address
- mfctr r0
- srawi r0, r0, 31 # apply expected sign to quotient
- xor r6, r6, r0 # RES = Q if CTR >= 0, -Q if CTR < 0
- xor r5, r5, r0
- subfc r4, r0, r6
- subfe r3, r0, r5
- blr
- .type __compcert_i64_sdiv, @function
- .size __compcert_i64_sdiv, .-__compcert_i64_sdiv
-
- \ No newline at end of file
diff --git a/runtime/powerpc/i64_shl.s b/runtime/powerpc/i64_shl.s
deleted file mode 100644
index f6edb6c2..00000000
--- a/runtime/powerpc/i64_shl.s
+++ /dev/null
@@ -1,64 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-# Shift left
-
- .balign 16
- .globl __compcert_i64_shl
-__compcert_i64_shl:
-# On PowerPC, shift instructions with amount mod 64 >= 32 return 0
-# hi = (hi << amount) | (lo >> (32 - amount)) | (lo << (amount - 32))
-# lo = lo << amount
-# if 0 <= amount < 32:
-# (amount - 32) mod 64 >= 32, hence lo << (amount - 32) == 0
-# if 32 <= amount < 64:
-# lo << amount == 0
-# (32 - amount) mod 64 >= 32, hence lo >> (32 - amount) == 0
- andi. r5, r5, 63 # take amount modulo 64
- subfic r6, r5, 32 # r6 = 32 - amount
- addi r7, r5, -32 # r7 = amount - 32
- slw r3, r3, r5
- srw r0, r4, r6
- or r3, r3, r0
- slw r0, r4, r7
- or r3, r3, r0
- slw r4, r4, r5
- blr
- .type __compcert_i64_shl, @function
- .size __compcert_i64_shl, .-__compcert_i64_shl
- \ No newline at end of file
diff --git a/runtime/powerpc/i64_shr.s b/runtime/powerpc/i64_shr.s
deleted file mode 100644
index b634aafd..00000000
--- a/runtime/powerpc/i64_shr.s
+++ /dev/null
@@ -1,65 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-# Shift right unsigned
-
- .balign 16
- .globl __compcert_i64_shr
-__compcert_i64_shr:
-# On PowerPC, shift instructions with amount mod 64 >= 32 return 0
-# lo = (lo >> amount) | (hi << (32 - amount)) | (hi >> (amount - 32))
-# hi = hi >> amount
-# if 0 <= amount < 32:
-# (amount - 32) mod 64 >= 32, hence hi >> (amount - 32) == 0
-# if 32 <= amount < 64:
-# hi >> amount == 0
-# (32 - amount) mod 64 >= 32, hence hi << (32 - amount) == 0
- andi. r5, r5, 63 # take amount modulo 64
- subfic r6, r5, 32 # r6 = 32 - amount
- addi r7, r5, -32 # r7 = amount - 32
- srw r4, r4, r5
- slw r0, r3, r6
- or r4, r4, r0
- srw r0, r3, r7
- or r4, r4, r0
- srw r3, r3, r5
- blr
- .type __compcert_i64_shr, @function
- .size __compcert_i64_shr, .-__compcert_i64_shr
-
- \ No newline at end of file
diff --git a/runtime/powerpc/i64_smod.s b/runtime/powerpc/i64_smod.s
deleted file mode 100644
index 6b4e1f89..00000000
--- a/runtime/powerpc/i64_smod.s
+++ /dev/null
@@ -1,70 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-## Signed remainder
-
- .balign 16
- .globl __compcert_i64_smod
-__compcert_i64_smod:
- mflr r0
- stw r0, 4(r1) # save return address in caller's frame
- mtctr r3 # save sign of result in CTR (sign of N)
- srawi r0, r3, 31 # take absolute value of N
- xor r4, r4, r0 # (i.e. N = N ^ r0 - r0,
- xor r3, r3, r0 # where r0 = 0 if N >= 0 and r0 = -1 if N < 0)
- subfc r4, r0, r4
- subfe r3, r0, r3
- srawi r0, r5, 31 # take absolute value of D
- xor r6, r6, r0 # (same trick)
- xor r5, r5, r0
- subfc r6, r0, r6
- subfe r5, r0, r5
- bl __compcert_i64_udivmod # do unsigned division
- lwz r0, 4(r1)
- mtlr r0 # restore return address
- mfctr r0
- srawi r0, r0, 31 # apply expected sign to remainder
- xor r4, r4, r0 # RES = R if CTR >= 0, -Q if CTR < 0
- xor r3, r3, r0
- subfc r4, r0, r4
- subfe r3, r0, r3
- blr
- .type __compcert_i64_smod, @function
- .size __compcert_i64_smod, .-__compcert_i64_smod
-
- \ No newline at end of file
diff --git a/runtime/powerpc/i64_smulh.s b/runtime/powerpc/i64_smulh.s
deleted file mode 100644
index 73393fce..00000000
--- a/runtime/powerpc/i64_smulh.s
+++ /dev/null
@@ -1,80 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris
-#
-# Copyright (c) 2016 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-### Signed multiply-high
-
-# Hacker's Delight section 8.3:
-# - compute high 64 bits of the unsigned product X * Y (see i64_umulh.S)
-# - subtract X if Y < 0
-# - subtract Y if X < 0
-
- .balign 16
- .globl __compcert_i64_smulh
-__compcert_i64_smulh:
-# r7:r8:r9 accumulate bits 127:32 of the full unsigned product
- mulhwu r9, r4, r6 # r9 = high half of XL.YL
- mullw r0, r4, r5 # r0 = low half of XL.YH
- addc r9, r9, r0
- mulhwu r0, r4, r5 # r0 = high half of XL.YH
- addze r8, r0
- mullw r0, r3, r6 # r0 = low half of XH.YL
- addc r9, r9, r0
- mulhwu r0, r3, r6 # r0 = high half of XH.YL
- adde r8, r8, r0
- li r7, 0
- addze r7, r7
- mullw r0, r3, r5 # r0 = low half of XH.YH
- addc r8, r8, r0
- mulhwu r0, r3, r5 # r0 = high half of XH.YH
- adde r7, r7, r0
-# Here r7:r8 contains the high 64 bits of the unsigned product.
-# Now, test signs and subtract if needed
- srawi r0, r3, 31 # r0 = -1 if X < 0, r0 = 0 if X >= 0
- srawi r9, r5, 31 # r9 = -1 if Y < 0, r9 = 0 if Y >= 0
- and r3, r3, r9 # set X = 0 if Y >= 0
- and r4, r4, r9
- and r5, r5, r0 # set Y = 0 if X >= 0
- and r6, r6, r0
- subfc r8, r4, r8 # subtract X
- subfe r7, r3, r7
- subfc r4, r6, r8 # subtract Y
- subfe r3, r5, r7
- blr
- .type __compcert_i64_smulh, @function
- .size __compcert_i64_smulh, .-__compcert_i64_smulh
-
diff --git a/runtime/powerpc/i64_stod.s b/runtime/powerpc/i64_stod.s
deleted file mode 100644
index 0c1ab720..00000000
--- a/runtime/powerpc/i64_stod.s
+++ /dev/null
@@ -1,67 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
-### Conversion from signed long to double float
-
- .balign 16
- .globl __compcert_i64_stod
-__compcert_i64_stod:
- addi r1, r1, -16
- lis r5, 0x4330
- li r6, 0
- stw r5, 0(r1)
- stw r4, 4(r1) # 0(r1) = 2^52 + (double) XL
- stw r5, 8(r1)
- stw r6, 12(r1) # 8(r1) = 2^52
- lfd f1, 0(r1)
- lfd f2, 8(r1)
- fsub f1, f1, f2 # f1 is XL (unsigned) as a double
- lis r5, 0x4530
- lis r6, 0x8000
- stw r5, 0(r1) # 0(r1) = 2^84 + ((double)XH - 2^31) * 2^32
- add r3, r3, r6
- stw r3, 4(r1)
- stw r5, 8(r1) # 8(r1) = 2^84 + 2^31 * 2^32
- stw r6, 12(r1)
- lfd f2, 0(r1)
- lfd f3, 8(r1)
- fsub f2, f2, f3 # f2 is XH (signed) * 2^32 as a double
- fadd f1, f1, f2 # add both to get result
- addi r1, r1, 16
- blr
- .type __compcert_i64_stod, @function
- .size __compcert_i64_stod, .-__compcert_i64_stod
-
diff --git a/runtime/powerpc/i64_stof.s b/runtime/powerpc/i64_stof.s
deleted file mode 100644
index 97fa6bb8..00000000
--- a/runtime/powerpc/i64_stof.s
+++ /dev/null
@@ -1,68 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-### Conversion from signed long to single float
-
- .balign 16
- .globl __compcert_i64_stof
-__compcert_i64_stof:
- mflr r9
- # Check whether -2^53 <= X < 2^53
- srawi r5, r3, 31
- srawi r6, r3, 21 # (r5,r6) = X >> 53
- addic r6, r6, 1
- addze r5, r5 # (r5,r6) = (X >> 53) + 1
- cmplwi r5, 2
- blt 1f
- # X is large enough that double rounding can occur.
- # Avoid it by nudging X away from the points where double rounding
- # occurs (the "round to odd" technique)
- rlwinm r0, r4, 0, 21, 31 # extract bits 0 to 11 of X
- addi r0, r0, 0x7FF # r0 = (X & 0x7FF) + 0x7FF
- # bit 12 of r0 is 0 if all low 12 bits of X are 0, 1 otherwise
- # bits 13-31 of r0 are 0
- or r4, r4, r0 # correct bit number 12 of X
- rlwinm r4, r4, 0, 0, 20 # set to 0 bits 0 to 11 of X
- # Convert to double, then round to single
-1: bl __compcert_i64_stod
- mtlr r9
- frsp f1, f1
- blr
- .type __compcert_i64_stof, @function
- .size __compcert_i64_stof, .-__compcert_i64_stof
-
diff --git a/runtime/powerpc/i64_udiv.s b/runtime/powerpc/i64_udiv.s
deleted file mode 100644
index e2da855a..00000000
--- a/runtime/powerpc/i64_udiv.s
+++ /dev/null
@@ -1,54 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-### Unsigned division
-
- .balign 16
- .globl __compcert_i64_udiv
-__compcert_i64_udiv:
- mflr r0
- stw r0, 4(r1) # save return address in caller's frame
- bl __compcert_i64_udivmod # unsigned divide
- lwz r0, 4(r1)
- mtlr r0 # restore return address
- mr r3, r5 # result = quotient
- mr r4, r6
- blr
- .type __compcert_i64_udiv, @function
- .size __compcert_i64_udiv, .-__compcert_i64_udiv
-
diff --git a/runtime/powerpc/i64_udivmod.s b/runtime/powerpc/i64_udivmod.s
deleted file mode 100644
index e81c6cef..00000000
--- a/runtime/powerpc/i64_udivmod.s
+++ /dev/null
@@ -1,234 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-# Unsigned division and modulus
-
-# This function computes both the quotient and the remainder of two
-# unsigned 64-bit integers.
-
-# Input: numerator N in (r3,r4), divisor D in (r5,r6)
-# Output: quotient Q in (r5,r6), remainder R in (r3,r4)
-# Destroys: all integer caller-save registers
-
- .globl __compcert_i64_udivmod
- .balign 16
-__compcert_i64_udivmod:
- cmplwi r5, 0 # DH == 0 ?
- stwu r1, -32(r1)
- mflr r0
- stw r0, 8(r1)
- stw r31, 12(r1)
- beq 1f
-# The general case
- stw r30, 16(r1)
- stw r29, 20(r1)
- stw r28, 24(r1)
- mr r28, r3 # Save N in (r28, r29)
- mr r29, r4
- mr r30, r5 # Save D in (r30, r31)
- mr r31, r6
- # Scale N and D down, giving N' and D', such that 2^31 <= D' < 2^32
- cntlzw r7, r5 # r7 = leading zeros in DH = 32 - shift amount
- subfic r8, r7, 32 # r8 = shift amount
- slw r0, r3, r7 # N' = N >> shift amount
- srw r3, r3, r8
- srw r4, r4, r8
- or r4, r4, r0
- slw r0, r5, r7 # D' = D >> shift amount
- srw r6, r6, r8
- or r5, r6, r0
- # Divide N' by D' to get an approximate quotient Q
- bl __compcert_i64_udiv6432 # r3 = quotient, r4 = remainder
- mr r6, r3 # low half of quotient Q
- li r5, 0 # high half of quotient is 0
- # Tentative quotient is either correct or one too high
- # Compute Q * D in (r7, r8)
-4: mullw r7, r6, r30 # r7 = Q * DH
- mullw r8, r6, r31 # r8 = low 32 bits of Q * DL
- mulhwu r0, r6, r31 # r0 = high 32 bits of Q * DL
- addc r7, r7, r0
- subfe. r0, r0, r0 # test carry: EQ iff carry
- beq 2f # handle overflow case
- # Compute R = N - Q * D, with borrow
- subfc r4, r8, r29
- subfe r3, r7, r28
- subfe. r0, r0, r0 # test borrow: EQ iff no borrow
- beq 3f # no borrow: N >= Q * D, we are good
- addi r6, r6, -1 # borrow: adjust Q down by 1
- addc r4, r4, r31 # and R up by D
- adde r3, r3, r30
- # Finished
-3: lwz r0, 8(r1)
- mtlr r0
- lwz r31, 12(r1)
- lwz r30, 16(r1)
- lwz r29, 20(r1)
- lwz r28, 24(r1)
- addi r1, r1, 32
- blr
- # Special case when Q * D overflows
-2: addi r6, r6, -1 # adjust Q down by 1
- b 4b # and redo computation and check of remainder
- .balign 16
-# Special case 64 bits divided by 32 bits
-1: cmplwi r3, 0 # NH == 0?
- beq 4f
- divwu r31, r3, r6 # Divide NH by DL, quotient QH in r31
- mullw r0, r31, r6
- subf r3, r0, r3 # NH is remainder of this division
- mr r5, r6
- bl __compcert_i64_udiv6432 # divide NH : NL by DL
- mr r5, r31 # high word of quotient
- mr r6, r3 # low word of quotient
- # r4 contains low word of remainder
- li r3, 0 # high word of remainder = 0
- lwz r0, 8(r1)
- mtlr r0
- lwz r31, 12(r1)
- addi r1, r1, 32
- blr
- .balign 16
-# Special case 32 bits divided by 32 bits
-4: mr r0, r6
- divwu r6, r4, r6 # low word of quotient
- li r5, 0 # high word of quotient is 0
- mullw r0, r6, r0
- subf r4, r0, r4 # low word of remainder
- li r3, 0 # high word of remainder is 0
- addi r1, r1, 32
- blr
-
- .type __compcert_i64_udivmod, @function
- .size __compcert_i64_udivmod, .-__compcert_i64_udivmod
-
-# Auxiliary division function: 64 bit integer divided by 32 bit integer
-# Not exported
-# Input: numerator N in (r3,r4), divisor D in r5
-# Output: quotient Q in r3, remainder R in r4
-# Destroys: all integer caller-save registers
-# Assumes: high word of N is less than D
-
- .balign 16
-__compcert_i64_udiv6432:
-# Algorithm 9.3 from Hacker's Delight, section 9.4
-# Initially: u1 in r3, u0 in r4, v in r5
-# s = __builtin_clz(v);
- cntlzw r6, r5 # s in r6
-# v = v << s;
- slw r5, r5, r6
-# vn1 = v >> 16; # vn1 in r7
- srwi r7, r5, 16
-# vn0 = v & 0xFFFF; # vn0 in r8
- rlwinm r8, r5, 0, 16, 31
-# un32 = (u1 << s) | (u0 >> 32 - s);
- subfic r0, r6, 32
- srw r0, r4, r0
- slw r3, r3, r6 # u1 dies, un32 in r3
- or r3, r3, r0
-# un10 = u0 << s;
- slw r4, r4, r6 # u0 dies, un10 in r4
-# un1 = un10 >> 16;
- srwi r9, r4, 16 # un1 in r9
-# un0 = un10 & 0xFFFF;
- rlwinm r4, r4, 0, 16, 31 # un10 dies, un0 in r4
-# q1 = un32/vn1;
- divwu r10, r3, r7 # q in r10
-# rhat = un32 - q1*vn1;
- mullw r0, r10, r7
- subf r11, r0, r3 # rhat in r11
-# again1:
-1:
-# if (q1 >= b || q1*vn0 > b*rhat + un1) {
- cmplwi r10, 0xFFFF
- bgt 2f
- mullw r0, r10, r8
- slwi r12, r11, 16
- add r12, r12, r9
- cmplw r0, r12
- ble 3f
-2:
-# q1 = q1 - 1;
- addi r10, r10, -1
-# rhat = rhat + vn1;
- add r11, r11, r7
-# if (rhat < b) goto again1;}
- cmplwi r11, 0xFFFF
- ble 1b
-3:
-# un21 = un32*b + un1 - q1*v;
- slwi r0, r3, 16 # un32 dies
- add r9, r0, r9 # un1 dies
- mullw r0, r10, r5
- subf r9, r0, r9 # un21 in r9
-# q0 = un21/vn1;
- divwu r3, r9, r7 # q0 in r3
-# rhat = un21 - q0*vn1;
- mullw r0, r3, r7
- subf r11, r0, r9 # rhat in r11
-# again2:
-4:
-# if (q0 >= b || q0*vn0 > b*rhat + un0) {
- cmplwi r3, 0xFFFF
- bgt 5f
- mullw r0, r3, r8
- slwi r12, r11, 16
- add r12, r12, r4
- cmplw r0, r12
- ble 6f
-5:
-# q0 = q0 - 1;
- addi r3, r3, -1
-# rhat = rhat + vn1;
- add r11, r11, r7
-# if (rhat < b) goto again2;}
- cmplwi r11, 0xFFFF
- ble 4b
-6:
-# remainder = (un21*b + un0 - q0*v) >> s;
- slwi r0, r9, 16
- add r4, r0, r4 # un0 dies, remainder in r4
- mullw r0, r3, r5
- subf r4, r0, r4
- srw r4, r4, r6
-# quotient = q1*b + q0;
- slwi r0, r10, 16
- add r3, r0, r3
- blr
-
- .type __compcert_i64_udiv6432, @function
- .size __compcert_i64_udiv6432,.-__compcert_i64_udiv6432
diff --git a/runtime/powerpc/i64_umod.s b/runtime/powerpc/i64_umod.s
deleted file mode 100644
index bf8d6121..00000000
--- a/runtime/powerpc/i64_umod.s
+++ /dev/null
@@ -1,47 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-### Unsigned modulus
-
- .balign 16
- .globl __compcert_i64_umod
-__compcert_i64_umod:
- b __compcert_i64_udivmod
- .type __compcert_i64_umod, @function
- .size __compcert_i64_umod, .-__compcert_i64_umod
-
diff --git a/runtime/powerpc/i64_umulh.s b/runtime/powerpc/i64_umulh.s
deleted file mode 100644
index 53b72948..00000000
--- a/runtime/powerpc/i64_umulh.s
+++ /dev/null
@@ -1,65 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris
-#
-# Copyright (c) 2016 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-### Unsigned multiply-high
-
-# X * Y = 2^64 XH.YH + 2^32 (XH.YL + XL.YH) + XL.YL
-
- .balign 16
- .globl __compcert_i64_umulh
-__compcert_i64_umulh:
-# r7:r8:r9 accumulate bits 127:32 of the full product
- mulhwu r9, r4, r6 # r9 = high half of XL.YL
- mullw r0, r4, r5 # r0 = low half of XL.YH
- addc r9, r9, r0
- mulhwu r0, r4, r5 # r0 = high half of XL.YH
- addze r8, r0
- mullw r0, r3, r6 # r0 = low half of XH.YL
- addc r9, r9, r0
- mulhwu r0, r3, r6 # r0 = high half of XH.YL
- adde r8, r8, r0
- li r7, 0
- addze r7, r7
- mullw r0, r3, r5 # r0 = low half of XH.YH
- addc r4, r8, r0
- mulhwu r0, r3, r5 # r0 = high half of XH.YH
- adde r3, r7, r0
- blr
- .type __compcert_i64_umulh, @function
- .size __compcert_i64_umulh, .-__compcert_i64_umulh
-
diff --git a/runtime/powerpc/i64_utod.s b/runtime/powerpc/i64_utod.s
deleted file mode 100644
index 69de6fdb..00000000
--- a/runtime/powerpc/i64_utod.s
+++ /dev/null
@@ -1,66 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-### Conversion from unsigned long to double float
-
- .balign 16
- .globl __compcert_i64_utod
-__compcert_i64_utod:
- addi r1, r1, -16
- lis r5, 0x4330
- li r6, 0
- stw r5, 0(r1)
- stw r4, 4(r1) # 0(r1) = 2^52 + (double) XL
- stw r5, 8(r1)
- stw r6, 12(r1) # 8(r1) = 2^52
- lfd f1, 0(r1)
- lfd f2, 8(r1)
- fsub f1, f1, f2 # f1 is (double) XL
- lis r5, 0x4530
- stw r5, 0(r1) # 0(r1) = 2^84 + (double) XH * 2^32
- stw r3, 4(r1)
- stw r5, 8(r1) # 8(r1) = 2^84
- lfd f2, 0(r1)
- lfd f3, 8(r1)
- fsub f2, f2, f3 # f2 is XH * 2^32 as a double
- fadd f1, f1, f2 # add both to get result
- addi r1, r1, 16
- blr
- .type __compcert_i64_utod, @function
- .size __compcert_i64_utod, .-__compcert_i64_utod
-
diff --git a/runtime/powerpc/i64_utof.s b/runtime/powerpc/i64_utof.s
deleted file mode 100644
index cdb2f867..00000000
--- a/runtime/powerpc/i64_utof.s
+++ /dev/null
@@ -1,64 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-### Conversion from unsigned long to single float
-
- .balign 16
- .globl __compcert_i64_utof
-__compcert_i64_utof:
- mflr r9
- # Check whether X < 2^53
- andis. r0, r3, 0xFFE0 # test bits 53...63 of X
- beq 1f
- # X is large enough that double rounding can occur.
- # Avoid it by nudging X away from the points where double rounding
- # occurs (the "round to odd" technique)
- rlwinm r0, r4, 0, 21, 31 # extract bits 0 to 11 of X
- addi r0, r0, 0x7FF # r0 = (X & 0x7FF) + 0x7FF
- # bit 12 of r0 is 0 if all low 12 bits of X are 0, 1 otherwise
- # bits 13-31 of r0 are 0
- or r4, r4, r0 # correct bit number 12 of X
- rlwinm r4, r4, 0, 0, 20 # set to 0 bits 0 to 11 of X
- # Convert to double, then round to single
-1: bl __compcert_i64_utod
- mtlr r9
- frsp f1, f1
- blr
- .type __compcert_i64_utof, @function
- .size __compcert_i64_utof, .-__compcert_i64_utof
-
diff --git a/runtime/powerpc/vararg.s b/runtime/powerpc/vararg.s
deleted file mode 100644
index 8d7e62c8..00000000
--- a/runtime/powerpc/vararg.s
+++ /dev/null
@@ -1,163 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for variadic functions <stdarg.h>. IA32 version
-
-# typedef struct {
-# unsigned char ireg; // index of next integer register
-# unsigned char freg; // index of next FP register
-# char * stk; // pointer to next argument in stack
-# struct {
-# int iregs[8];
-# double fregs[8];
-# } * regs; // pointer to saved register area
-# } va_list[1];
-#
-# unsigned int __compcert_va_int32(va_list ap);
-# unsigned long long __compcert_va_int64(va_list ap);
-# double __compcert_va_float64(va_list ap);
-
- .text
-
- .balign 16
- .globl __compcert_va_int32
-__compcert_va_int32:
- # r3 = ap = address of va_list structure
- lbz r4, 0(r3) # r4 = ap->ireg = next integer register
- cmplwi r4, 8
- bge 1f
- # Next argument was passed in an integer register
- lwz r5, 8(r3) # r5 = ap->regs = base of saved register area
- rlwinm r6, r4, 2, 0, 29 # r6 = r4 * 4
- addi r4, r4, 1 # increment ap->ireg
- stb r4, 0(r3)
- lwzx r3, r5, r6 # load argument in r3
- blr
- # Next argument was passed on stack
-1: lwz r5, 4(r3) # r5 = ap->stk = next argument passed on stack
- addi r5, r5, 4 # advance ap->stk by 4
- stw r5, 4(r3)
- lwz r3, -4(r5) # load argument in r3
- blr
- .type __compcert_va_int32, @function
- .size __compcert_va_int32, .-__compcert_va_int32
-
- .balign 16
- .globl __compcert_va_int64
-__compcert_va_int64:
- # r3 = ap = address of va_list structure
- lbz r4, 0(r3) # r4 = ap->ireg = next integer register
- cmplwi r4, 7
- bge 1f
- # Next argument was passed in two consecutive integer register
- lwz r5, 8(r3) # r5 = ap->regs = base of saved register area
- addi r4, r4, 3 # round r4 up to an even number and add 2
- rlwinm r4, r4, 0, 0, 30
- rlwinm r6, r4, 2, 0, 29 # r6 = r4 * 4
- add r5, r5, r6 # r5 = address of argument + 8
- stb r4, 0(r3) # update ap->ireg
- lwz r3, -8(r5) # load argument in r3:r4
- lwz r4, -4(r5)
- blr
- # Next argument was passed on stack
-1: lwz r5, 4(r3) # r5 = ap->stk = next argument passed on stack
- li r4, 8
- stb r4, 0(r3) # set ap->ireg = 8 so that no ireg is left
- addi r5, r5, 15 # round r5 to a multiple of 8 and add 8
- rlwinm r5, r5, 0, 0, 28
- stw r5, 4(r3) # update ap->stk
- lwz r3, -8(r5) # load argument in r3:r4
- lwz r4, -4(r5)
- blr
- .type __compcert_va_int64, @function
- .size __compcert_va_int64, .-__compcert_va_int64
-
- .balign 16
- .globl __compcert_va_float64
-__compcert_va_float64:
- # r3 = ap = address of va_list structure
- lbz r4, 1(r3) # r4 = ap->freg = next float register
- cmplwi r4, 8
- bge 1f
- # Next argument was passed in a FP register
- lwz r5, 8(r3) # r5 = ap->regs = base of saved register area
- rlwinm r6, r4, 3, 0, 28 # r6 = r4 * 8
- add r5, r5, r6
- lfd f1, 32(r5) # load argument in f1
- addi r4, r4, 1 # increment ap->freg
- stb r4, 1(r3)
- blr
- # Next argument was passed on stack
-1: lwz r5, 4(r3) # r5 = ap->stk = next argument passed on stack
- addi r5, r5, 15 # round r5 to a multiple of 8 and add 8
- rlwinm r5, r5, 0, 0, 28
- lfd f1, -8(r5) # load argument in f1
- stw r5, 4(r3) # update ap->stk
- blr
- .type __compcert_va_float64, @function
- .size __compcert_va_float64, .-__compcert_va_int64
-
- .balign 16
- .globl __compcert_va_composite
-__compcert_va_composite:
- b __compcert_va_int32
- .type __compcert_va_composite, @function
- .size __compcert_va_composite, .-__compcert_va_composite
-
-# Save integer and FP registers at beginning of vararg function
-
- .balign 16
- .globl __compcert_va_saveregs
-__compcert_va_saveregs:
- lwz r11, 0(r1) # r11 point to top of our frame
- stwu r3, -96(r11) # register save area is 96 bytes below
- stw r4, 4(r11)
- stw r5, 8(r11)
- stw r6, 12(r11)
- stw r7, 16(r11)
- stw r8, 20(r11)
- stw r9, 24(r11)
- stw r10, 28(r11)
- bf 6, 1f # don't save FP regs if CR6 bit is clear
- stfd f1, 32(r11)
- stfd f2, 40(r11)
- stfd f3, 48(r11)
- stfd f4, 56(r11)
- stfd f5, 64(r11)
- stfd f6, 72(r11)
- stfd f7, 80(r11)
- stfd f8, 88(r11)
-1: blr
- .type __compcert_va_saveregs, @function
- .size __compcert_va_saveregs, .-__compcert_va_saveregs
diff --git a/runtime/powerpc64/i64_dtou.s b/runtime/powerpc64/i64_dtou.s
deleted file mode 100644
index e58bcfaf..00000000
--- a/runtime/powerpc64/i64_dtou.s
+++ /dev/null
@@ -1,66 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC 64 version.
-
- .text
-
-### Conversion from double float to unsigned long
-
- .balign 16
- .globl __compcert_i64_dtou
-__compcert_i64_dtou:
- lis r0, 0x5f00 # 0x5f00_0000 = 2^63 in binary32 format
- stwu r0, -16(r1)
- lfs f2, 0(r1) # f2 = 2^63
- fcmpu cr0, f1, f2 # crbit 0 is f1 < f2
- bf 0, 1f # branch if f1 >= 2^63 (or f1 is NaN)
- fctidz f1, f1 # convert as signed
- stfd f1, 0(r1)
- lwz r3, 0(r1)
- lwz r4, 4(r1)
- addi r1, r1, 16
- blr
-1: fsub f1, f1, f2 # shift argument down by 2^63
- fctidz f1, f1 # convert as signed
- stfd f1, 0(r1)
- lwz r3, 0(r1)
- lwz r4, 4(r1)
- addis r3, r3, 0x8000 # shift result up by 2^63
- addi r1, r1, 16
- blr
- .type __compcert_i64_dtou, @function
- .size __compcert_i64_dtou, .-__compcert_i64_dtou
-
-
diff --git a/runtime/powerpc64/i64_stof.s b/runtime/powerpc64/i64_stof.s
deleted file mode 100644
index 779cbc18..00000000
--- a/runtime/powerpc64/i64_stof.s
+++ /dev/null
@@ -1,68 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC 64 version.
-
- .text
-
-### Conversion from signed long to single float
-
- .balign 16
- .globl __compcert_i64_stof
-__compcert_i64_stof:
- rldimi r4, r3, 32, 0 # reassemble (r3,r4) as a 64-bit integer in r4
- # Check whether -2^53 <= X < 2^53
- sradi r5, r4, 53
- addi r5, r5, 1
- cmpldi r5, 2
- blt 1f
- # X is large enough that double rounding can occur.
- # Avoid it by nudging X away from the points where double rounding
- # occurs (the "round to odd" technique)
- rldicl r5, r4, 0, 53 # extract bits 0 to 11 of X
- addi r5, r5, 0x7FF # r5 = (X & 0x7FF) + 0x7FF
- # bit 12 of r5 is 0 if all low 12 bits of X are 0, 1 otherwise
- # bits 13-63 of r5 are 0
- or r4, r4, r5 # correct bit number 12 of X
- rldicr r4, r4, 0, 52 # set to 0 bits 0 to 11 of X
- # Convert to double, then round to single
-1: stdu r4, -16(r1)
- lfd f1, 0(r1)
- fcfid f1, f1
- frsp f1, f1
- addi r1, r1, 16
- blr
- .type __compcert_i64_stof, @function
- .size __compcert_i64_stof, .-__compcert_i64_stof
-
diff --git a/runtime/powerpc64/i64_utod.s b/runtime/powerpc64/i64_utod.s
deleted file mode 100644
index 491ee26b..00000000
--- a/runtime/powerpc64/i64_utod.s
+++ /dev/null
@@ -1,79 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC 64 version.
-
- .text
-
-### Conversion from unsigned long to double float
-
- .balign 16
- .globl __compcert_i64_utod
-__compcert_i64_utod:
- rldicl r3, r3, 0, 32 # clear top 32 bits
- rldicl r4, r4, 0, 32 # clear top 32 bits
- lis r5, 0x4f80 # 0x4f80_0000 = 2^32 in binary32 format
- stdu r3, -32(r1)
- std r4, 8(r1)
- stw r5, 16(r1)
- lfd f1, 0(r1) # high 32 bits of argument
- lfd f2, 8(r1) # low 32 bits of argument
- lfs f3, 16(r1) # 2^32
- fcfid f1, f1 # convert both 32-bit halves to FP (exactly)
- fcfid f2, f2
- fmadd f1, f1, f3, f2 # compute hi * 2^32 + lo
- addi r1, r1, 32
- blr
- .type __compcert_i64_utod, @function
- .size __compcert_i64_utod, .-__compcert_i64_utod
-
-# Alternate implementation using round-to-odd:
-# rldimi r4, r3, 32, 0 # reassemble (r3,r4) as a 64-bit integer in r4
-# cmpdi r4, 0 # is r4 >= 2^63 ?
-# blt 1f
-# stdu r4, -16(r1) # r4 < 2^63: convert as signed
-# lfd f1, 0(r1)
-# fcfid f1, f1
-# addi r1, r1, 16
-# blr
-#1: rldicl r0, r4, 0, 63 # extract low bit of r4
-# srdi r4, r4, 1
-# or r4, r4, r0 # round r4 to 63 bits, using round-to-odd
-# stdu r4, -16(r1) # convert to binary64
-# lfd f1, 0(r1)
-# fcfid f1, f1
-# fadd f1, f1, f1 # multiply result by 2
-# addi r1, r1, 16
-# blr
- \ No newline at end of file
diff --git a/runtime/powerpc64/i64_utof.s b/runtime/powerpc64/i64_utof.s
deleted file mode 100644
index cdb2f867..00000000
--- a/runtime/powerpc64/i64_utof.s
+++ /dev/null
@@ -1,64 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for 64-bit integer arithmetic. PowerPC version.
-
- .text
-
-### Conversion from unsigned long to single float
-
- .balign 16
- .globl __compcert_i64_utof
-__compcert_i64_utof:
- mflr r9
- # Check whether X < 2^53
- andis. r0, r3, 0xFFE0 # test bits 53...63 of X
- beq 1f
- # X is large enough that double rounding can occur.
- # Avoid it by nudging X away from the points where double rounding
- # occurs (the "round to odd" technique)
- rlwinm r0, r4, 0, 21, 31 # extract bits 0 to 11 of X
- addi r0, r0, 0x7FF # r0 = (X & 0x7FF) + 0x7FF
- # bit 12 of r0 is 0 if all low 12 bits of X are 0, 1 otherwise
- # bits 13-31 of r0 are 0
- or r4, r4, r0 # correct bit number 12 of X
- rlwinm r4, r4, 0, 0, 20 # set to 0 bits 0 to 11 of X
- # Convert to double, then round to single
-1: bl __compcert_i64_utod
- mtlr r9
- frsp f1, f1
- blr
- .type __compcert_i64_utof, @function
- .size __compcert_i64_utof, .-__compcert_i64_utof
-
diff --git a/runtime/powerpc64/vararg.s b/runtime/powerpc64/vararg.s
deleted file mode 100644
index 8d7e62c8..00000000
--- a/runtime/powerpc64/vararg.s
+++ /dev/null
@@ -1,163 +0,0 @@
-# *****************************************************************
-#
-# The Compcert verified compiler
-#
-# Xavier Leroy, INRIA Paris-Rocquencourt
-#
-# Copyright (c) 2013 Institut National de Recherche en Informatique et
-# en Automatique.
-#
-# Redistribution and use in source and binary forms, with or without
-# modification, are permitted provided that the following conditions are met:
-# * Redistributions of source code must retain the above copyright
-# notice, this list of conditions and the following disclaimer.
-# * Redistributions in binary form must reproduce the above copyright
-# notice, this list of conditions and the following disclaimer in the
-# documentation and/or other materials provided with the distribution.
-# * Neither the name of the <organization> nor the
-# names of its contributors may be used to endorse or promote products
-# derived from this software without specific prior written permission.
-#
-# THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
-# "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
-# LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
-# A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
-# HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
-# EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
-# PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
-# PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
-# LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
-# NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
-# SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
-#
-# *********************************************************************
-
-# Helper functions for variadic functions <stdarg.h>. IA32 version
-
-# typedef struct {
-# unsigned char ireg; // index of next integer register
-# unsigned char freg; // index of next FP register
-# char * stk; // pointer to next argument in stack
-# struct {
-# int iregs[8];
-# double fregs[8];
-# } * regs; // pointer to saved register area
-# } va_list[1];
-#
-# unsigned int __compcert_va_int32(va_list ap);
-# unsigned long long __compcert_va_int64(va_list ap);
-# double __compcert_va_float64(va_list ap);
-
- .text
-
- .balign 16
- .globl __compcert_va_int32
-__compcert_va_int32:
- # r3 = ap = address of va_list structure
- lbz r4, 0(r3) # r4 = ap->ireg = next integer register
- cmplwi r4, 8
- bge 1f
- # Next argument was passed in an integer register
- lwz r5, 8(r3) # r5 = ap->regs = base of saved register area
- rlwinm r6, r4, 2, 0, 29 # r6 = r4 * 4
- addi r4, r4, 1 # increment ap->ireg
- stb r4, 0(r3)
- lwzx r3, r5, r6 # load argument in r3
- blr
- # Next argument was passed on stack
-1: lwz r5, 4(r3) # r5 = ap->stk = next argument passed on stack
- addi r5, r5, 4 # advance ap->stk by 4
- stw r5, 4(r3)
- lwz r3, -4(r5) # load argument in r3
- blr
- .type __compcert_va_int32, @function
- .size __compcert_va_int32, .-__compcert_va_int32
-
- .balign 16
- .globl __compcert_va_int64
-__compcert_va_int64:
- # r3 = ap = address of va_list structure
- lbz r4, 0(r3) # r4 = ap->ireg = next integer register
- cmplwi r4, 7
- bge 1f
- # Next argument was passed in two consecutive integer register
- lwz r5, 8(r3) # r5 = ap->regs = base of saved register area
- addi r4, r4, 3 # round r4 up to an even number and add 2
- rlwinm r4, r4, 0, 0, 30
- rlwinm r6, r4, 2, 0, 29 # r6 = r4 * 4
- add r5, r5, r6 # r5 = address of argument + 8
- stb r4, 0(r3) # update ap->ireg
- lwz r3, -8(r5) # load argument in r3:r4
- lwz r4, -4(r5)
- blr
- # Next argument was passed on stack
-1: lwz r5, 4(r3) # r5 = ap->stk = next argument passed on stack
- li r4, 8
- stb r4, 0(r3) # set ap->ireg = 8 so that no ireg is left
- addi r5, r5, 15 # round r5 to a multiple of 8 and add 8
- rlwinm r5, r5, 0, 0, 28
- stw r5, 4(r3) # update ap->stk
- lwz r3, -8(r5) # load argument in r3:r4
- lwz r4, -4(r5)
- blr
- .type __compcert_va_int64, @function
- .size __compcert_va_int64, .-__compcert_va_int64
-
- .balign 16
- .globl __compcert_va_float64
-__compcert_va_float64:
- # r3 = ap = address of va_list structure
- lbz r4, 1(r3) # r4 = ap->freg = next float register
- cmplwi r4, 8
- bge 1f
- # Next argument was passed in a FP register
- lwz r5, 8(r3) # r5 = ap->regs = base of saved register area
- rlwinm r6, r4, 3, 0, 28 # r6 = r4 * 8
- add r5, r5, r6
- lfd f1, 32(r5) # load argument in f1
- addi r4, r4, 1 # increment ap->freg
- stb r4, 1(r3)
- blr
- # Next argument was passed on stack
-1: lwz r5, 4(r3) # r5 = ap->stk = next argument passed on stack
- addi r5, r5, 15 # round r5 to a multiple of 8 and add 8
- rlwinm r5, r5, 0, 0, 28
- lfd f1, -8(r5) # load argument in f1
- stw r5, 4(r3) # update ap->stk
- blr
- .type __compcert_va_float64, @function
- .size __compcert_va_float64, .-__compcert_va_int64
-
- .balign 16
- .globl __compcert_va_composite
-__compcert_va_composite:
- b __compcert_va_int32
- .type __compcert_va_composite, @function
- .size __compcert_va_composite, .-__compcert_va_composite
-
-# Save integer and FP registers at beginning of vararg function
-
- .balign 16
- .globl __compcert_va_saveregs
-__compcert_va_saveregs:
- lwz r11, 0(r1) # r11 point to top of our frame
- stwu r3, -96(r11) # register save area is 96 bytes below
- stw r4, 4(r11)
- stw r5, 8(r11)
- stw r6, 12(r11)
- stw r7, 16(r11)
- stw r8, 20(r11)
- stw r9, 24(r11)
- stw r10, 28(r11)
- bf 6, 1f # don't save FP regs if CR6 bit is clear
- stfd f1, 32(r11)
- stfd f2, 40(r11)
- stfd f3, 48(r11)
- stfd f4, 56(r11)
- stfd f5, 64(r11)
- stfd f6, 72(r11)
- stfd f7, 80(r11)
- stfd f8, 88(r11)
-1: blr
- .type __compcert_va_saveregs, @function
- .size __compcert_va_saveregs, .-__compcert_va_saveregs
diff --git a/test/monniaux/BearSSL/mk/mkT0.cmd b/test/monniaux/BearSSL/mk/mkT0.cmd
deleted file mode 100644
index 98955625..00000000
--- a/test/monniaux/BearSSL/mk/mkT0.cmd
+++ /dev/null
@@ -1,32 +0,0 @@
-@echo off
-
-rem =====================================================================
-rem This script uses the command-line C# compiler csc.exe, which is
-rem provided with the .NET framework. We need framework 3.5 or later
-rem (some of the code uses features not available in the language version
-rem implemented in the compiler provided with framework 2.0.50727).
-rem =====================================================================
-
-if exist "%SystemRoot%\Microsoft.NET\Framework\v3.5\csc.exe" (
- set CSC="%SystemRoot%\Microsoft.NET\Framework\v3.5\csc.exe"
- goto do_compile
-)
-if exist "%SystemRoot%\Microsoft.NET\Framework\v4.0.30319\csc.exe" (
- set CSC="%SystemRoot%\Microsoft.NET\Framework\v4.0.30319\csc.exe"
- goto do_compile
-)
-if exist "%SystemRoot%\Microsoft.NET\Framework64\v3.5\csc.exe" (
- set CSC="%SystemRoot%\Microsoft.NET\Framework64\v3.5\csc.exe"
- goto do_compile
-)
-if exist "%SystemRoot%\Microsoft.NET\Framework64\v4.0.30319\csc.exe" (
- set CSC="%SystemRoot%\Microsoft.NET\Framework64\v4.0.30319\csc.exe"
- goto do_compile
-)
-
-echo C# compiler not found
-exit 1
-
-:do_compile
-%CSC% /nologo /out:T0Comp.exe /main:T0Comp /res:T0\kern.t0,t0-kernel T0\*.cs
-if %errorlevel% neq 0 exit /b %errorlevel%
diff --git a/test/monniaux/builtins/fma.c b/test/monniaux/builtins/fma.c
new file mode 100644
index 00000000..4083e781
--- /dev/null
+++ b/test/monniaux/builtins/fma.c
@@ -0,0 +1,14 @@
+#define COMPCERT_NO_FP_MACROS
+#include <math.h>
+#include <stdio.h>
+#include <stdlib.h>
+
+int main(int argc, char **argv) {
+ if (argc < 4) return 1;
+ double x = strtod(argv[1], NULL);
+ double y = strtod(argv[2], NULL);
+ double z = strtod(argv[3], NULL);
+ printf("%g %g %g\n", __builtin_fma(x, y, z), fma(x, y, z), x*y + z);
+ printf("%g %g %g\n", __builtin_fma(-x, y, z), fma(-x, y, z), (-x)*y + z);
+ return 0;
+}
diff --git a/test/monniaux/latency/latency.s b/test/monniaux/latency/latency.s
deleted file mode 100644
index 494f9f77..00000000
--- a/test/monniaux/latency/latency.s
+++ /dev/null
@@ -1,27 +0,0 @@
-# File generated by CompCert 3.4
-# Command line: -S latency.c
- .text
- .balign 2
- .globl latency
-latency:
- make $r2 = 100
-;;
- loopdo $r2, .exitloop
-;;
- lws $r1 = 0[$r0]
-;;
- addw $r0 = $r0, 0
-;;
- addw $r0 = $r0, 0
-;;
- addw $r0 = $r0, 0
-;;
- addw $r1 = $r1, 1
-;;
- sw 0[$r0] = $r1
-;;
-.exitloop:
- ret
-;;
- .type latency, @function
- .size latency, . - latency
diff --git a/test/monniaux/predicated/predicated.s b/test/monniaux/predicated/predicated.s
deleted file mode 100644
index f11606c2..00000000
--- a/test/monniaux/predicated/predicated.s
+++ /dev/null
@@ -1,13 +0,0 @@
- .text
-
- .globl predicated_write
-predicated_write:
- sd.wnez $r0? 8[$r1] = $r2
- ret
- ;;
-
- .globl predicated_read
-predicated_read:
- ld.wnez $r1? $r0 = 8[$r2]
- ret
- ;;
diff --git a/test/mppa/instr/builtin32.c b/test/mppa/instr/builtin32.c
new file mode 100644
index 00000000..c7689dc8
--- /dev/null
+++ b/test/mppa/instr/builtin32.c
@@ -0,0 +1,12 @@
+#include "framework.h"
+
+BEGIN_TEST(int)
+ int *ptr = &c;
+#ifdef __K1C__
+ int d = c;
+ a = __builtin_k1_alclrw(ptr);
+ c = d;
+
+#endif
+END_TEST32()
+
diff --git a/test/regression/builtins-arm.c b/test/regression/builtins-arm.c
index 709343ce..d06e8e5e 100644
--- a/test/regression/builtins-arm.c
+++ b/test/regression/builtins-arm.c
@@ -2,14 +2,15 @@
#include <stdio.h>
+unsigned int x = 0x12345678;
+unsigned int y = 0xDEADBEEF;
+unsigned long long xx = 0x1234567812345678ULL;
+double a = 3.14159;
+unsigned short s = 0x1234;
+
int main(int argc, char ** argv)
{
- unsigned int x = 0x12345678;
- unsigned int y = 0xDEADBEEF;
- unsigned long long xx = 0x1234567812345678ULL;
unsigned z;
- double a = 3.14159;
- unsigned short s = 0x1234;
printf("bswap(%x) = %x\n", x, __builtin_bswap(x));
printf("bswap16(%x) = %x\n", s, __builtin_bswap16(s));
diff --git a/test/regression/builtins-powerpc.c b/test/regression/builtins-powerpc.c
index 23e9d191..8fd5818b 100644
--- a/test/regression/builtins-powerpc.c
+++ b/test/regression/builtins-powerpc.c
@@ -9,16 +9,17 @@ char * check_relative_error(double exact, double actual, double precision)
return fabs(relative_error) <= precision ? "OK" : "ERROR";
}
+unsigned int x = 0x12345678;
+unsigned int y = 0xDEADBEEF;
+unsigned long long xx = 0x1234567812345678ULL;
+double a = 3.14159;
+double b = 2.718;
+double c = 1.414;
+unsigned short s = 0x1234;
+
int main(int argc, char ** argv)
{
- unsigned int x = 0x12345678;
- unsigned int y = 0xDEADBEEF;
- unsigned long long xx = 0x1234567812345678ULL;
unsigned z;
- double a = 3.14159;
- double b = 2.718;
- double c = 1.414;
- unsigned short s = 0x1234;
printf("mulhw(%x, %x) = %x\n", x, y, __builtin_mulhw(x, y));
printf("mulhwu(%x, %x) = %x\n", x, y, __builtin_mulhwu(x, y));
diff --git a/test/regression/builtins-riscV.c b/test/regression/builtins-riscV.c
index a302a6c4..c34fdf2c 100644
--- a/test/regression/builtins-riscV.c
+++ b/test/regression/builtins-riscV.c
@@ -2,15 +2,15 @@
#include <stdio.h>
+unsigned int x = 0x12345678;
+unsigned short s = 0x1234;
+unsigned long long zz = 0x123456789ABCDEF0ULL;
+double a = 3.14159;
+double b = 2.718;
+double c = 1.414;
+
int main(int argc, char ** argv)
{
- unsigned int x = 0x12345678;
- unsigned short s = 0x1234;
- unsigned long long zz = 0x123456789ABCDEF0ULL;
- double a = 3.14159;
- double b = 2.718;
- double c = 1.414;
-
printf("bswap16(%x) = %x\n", s, __builtin_bswap16(s));
printf("bswap32(%x) = %x\n", x, __builtin_bswap32(x));
printf("bswap64(%llx) = %llx\n", zz, __builtin_bswap64(zz));
diff --git a/test/regression/builtins-x86.c b/test/regression/builtins-x86.c
index 1ba213e7..6233f9fd 100644
--- a/test/regression/builtins-x86.c
+++ b/test/regression/builtins-x86.c
@@ -2,18 +2,19 @@
#include <stdio.h>
+unsigned int x = 0x12345678;
+unsigned int y = 0xDEADBEEF;
+unsigned long long xx = 0x1234567812345678ULL;
+unsigned long long yy = 0x1234567800000000ULL;
+unsigned long long zz = 0x123456789ABCDEF0ULL;
+double a = 3.14159;
+double b = 2.718;
+double c = 1.414;
+unsigned short s = 0x1234;
+
int main(int argc, char ** argv)
{
- unsigned int x = 0x12345678;
- unsigned int y = 0xDEADBEEF;
- unsigned long long xx = 0x1234567812345678ULL;
- unsigned long long yy = 0x1234567800000000ULL;
- unsigned long long zz = 0x123456789ABCDEF0ULL;
unsigned z;
- double a = 3.14159;
- double b = 2.718;
- double c = 1.414;
- unsigned short s = 0x1234;
printf("bswap(%x) = %x\n", x, __builtin_bswap(x));
printf("bswap16(%x) = %x\n", s, __builtin_bswap16(s));
diff --git a/test/regression/extasm.c b/test/regression/extasm.c
index babc57f1..83a07a05 100644
--- a/test/regression/extasm.c
+++ b/test/regression/extasm.c
@@ -33,6 +33,7 @@ int main()
void * y;
long long z;
double f;
+ float sf;
char c[16];
/* No inputs, no outputs */
@@ -72,6 +73,15 @@ int main()
#ifdef FAILURES
asm("FAIL4 a:%[a]" : "=r"(x) : [z]"i"(0));
#endif
+ /* One argument of each type */
+ asm("TEST15 int32 %0" : : "r" (x));
+#ifdef SIXTYFOUR
+ asm("TEST15 int64 %0" : : "r" (z));
+#else
+ asm("TEST15 int64 %Q0 / %R0" : : "r" (z));
+#endif
+ asm("TEST15 float64 %0" : : "r" (f));
+ asm("TEST15 float32 %0" : : "r" (sf));
/* Various failures */
#ifdef FAILURES
asm("FAIL5 out:%0,%1" : "=r"(x), "=r"(y));
diff --git a/test/regression/ifconv.c b/test/regression/ifconv.c
index dcbf43e5..e12a394c 100644
--- a/test/regression/ifconv.c
+++ b/test/regression/ifconv.c
@@ -83,6 +83,26 @@ float sdoz(float x, float y)
return x >= y ? x - y : 0.0f;
}
+/* Examples where constant propagation should take place */
+
+int constprop1(int x)
+{
+ int n = 0;
+ return n ? x : 42;
+}
+
+int constprop2(int x)
+{
+ int n = 1;
+ return n ? x : 42;
+}
+
+int constprop3(int x, int y)
+{
+ int n = 0;
+ return x < n ? y - 1 : y + 1;
+}
+
/* Test harness */
#define TESTI(call) printf(#call " = %d\n", call)
diff --git a/x86/Builtins1.v b/x86/Builtins1.v
new file mode 100644
index 00000000..6103cc4c
--- /dev/null
+++ b/x86/Builtins1.v
@@ -0,0 +1,54 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, Collège de France and Inria Paris *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the GNU General Public License as published by *)
+(* the Free Software Foundation, either version 2 of the License, or *)
+(* (at your option) any later version. This file is also distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Platform-specific built-in functions *)
+
+Require Import String Coqlib.
+Require Import AST Integers Floats Values.
+Require Import Builtins0.
+
+Inductive platform_builtin : Type :=
+ | BI_fmin
+ | BI_fmax.
+
+Local Open Scope string_scope.
+
+Definition platform_builtin_table : list (string * platform_builtin) :=
+ ("__builtin_fmin", BI_fmin)
+ :: ("__builtin_fmax", BI_fmax)
+ :: nil.
+
+Definition platform_builtin_sig (b: platform_builtin) : signature :=
+ match b with
+ | BI_fmin | BI_fmax =>
+ mksignature (Tfloat :: Tfloat :: nil) (Some Tfloat) cc_default
+ end.
+
+Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (proj_sig_res (platform_builtin_sig b)) :=
+ match b with
+ | BI_fmin =>
+ mkbuiltin_n2t Tfloat Tfloat Tfloat
+ (fun f1 f2 => match Float.compare f1 f2 with
+ | Some Eq | Some Lt => f1
+ | Some Gt | None => f2
+ end)
+ | BI_fmax =>
+ mkbuiltin_n2t Tfloat Tfloat Tfloat
+ (fun f1 f2 => match Float.compare f1 f2 with
+ | Some Eq | Some Gt => f1
+ | Some Lt | None => f2
+ end)
+ end.
+
diff --git a/x86/CBuiltins.ml b/x86/CBuiltins.ml
index 69a2eb64..f4f40a31 100644
--- a/x86/CBuiltins.ml
+++ b/x86/CBuiltins.ml
@@ -26,13 +26,11 @@ let (va_list_type, va_list_scalar, size_va_list) =
(TPtr(TVoid [], []), true, 4)
let builtins = {
- Builtins.typedefs = [
+ builtin_typedefs = [
"__builtin_va_list", va_list_type;
];
- Builtins.functions = [
+ builtin_functions = [
(* Integer arithmetic *)
- "__builtin_bswap64",
- (TInt(IULongLong, []), [TInt(IULongLong, [])], false);
"__builtin_clz",
(TInt(IInt, []), [TInt(IUInt, [])], false);
"__builtin_clzl",
diff --git a/x86/ConstpropOp.vp b/x86/ConstpropOp.vp
index f59b9dba..ada8d54a 100644
--- a/x86/ConstpropOp.vp
+++ b/x86/ConstpropOp.vp
@@ -16,7 +16,7 @@
Require Import Coqlib Compopts.
Require Import AST Integers Floats.
Require Import Op Registers.
-Require Import ValueDomain.
+Require Import ValueDomain ValueAOp.
(** * Converting known values to constants *)
@@ -98,6 +98,15 @@ Nondetfunction make_cmp (c: condition) (args: list reg) (vl: list aval) :=
make_cmp_base c args vl
end.
+Definition make_select (c: condition) (ty: typ)
+ (r1 r2: reg) (args: list reg) (vl: list aval) :=
+ match resolve_branch (eval_static_condition c vl) with
+ | Some b => (Omove, (if b then r1 else r2) :: nil)
+ | None =>
+ let (c', args') := cond_strength_reduction c args vl in
+ (Osel c' ty, r1 :: r2 :: args')
+ end.
+
(** For addressing modes, we need to distinguish
- reductions that produce pointers (i.e. that produce [Aglobal], [Ainstack], [Abased] and [Abasedscaled] addressing modes), which are valid only if the pointer size is right;
- other reductions (producing [Aindexed] or [Aindexed2] modes), which are valid independently of the pointer size.
@@ -416,6 +425,7 @@ Nondetfunction op_strength_reduction
let (addr', args') := addr_strength_reduction_64 addr args vl in
(Oleal addr', args')
| Ocmp c, args, vl => make_cmp c args vl
+ | Osel c ty, r1 :: r2 :: args, v1 :: v2 :: vl => make_select c ty r1 r2 args vl
| Omulf, r1 :: r2 :: nil, v1 :: F n2 :: nil => make_mulfimm n2 r1 r1 r2
| Omulf, r1 :: r2 :: nil, F n1 :: v2 :: nil => make_mulfimm n1 r2 r1 r2
| Omulfs, r1 :: r2 :: nil, v1 :: FS n2 :: nil => make_mulfsimm n2 r1 r1 r2
diff --git a/x86/ConstpropOpproof.v b/x86/ConstpropOpproof.v
index 3bb0f3cd..6d2df9c1 100644
--- a/x86/ConstpropOpproof.v
+++ b/x86/ConstpropOpproof.v
@@ -14,7 +14,7 @@
Require Import Coqlib Compopts.
Require Import Integers Floats Values Memory Globalenvs Events.
-Require Import Op Registers RTL ValueDomain.
+Require Import Op Registers RTL ValueDomain ValueAOp ValueAnalysis.
Require Import ConstpropOp.
Section STRENGTH_REDUCTION.
@@ -371,6 +371,28 @@ Proof.
- apply make_cmp_base_correct; auto.
Qed.
+Lemma make_select_correct:
+ forall c ty r1 r2 args vl,
+ vl = map (fun r => AE.get r ae) args ->
+ let (op', args') := make_select c ty r1 r2 args vl in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op' e##args' m = Some v
+ /\ Val.lessdef (Val.select (eval_condition c e##args m) e#r1 e#r2 ty) v.
+Proof.
+ unfold make_select; intros.
+ destruct (resolve_branch (eval_static_condition c vl)) as [b|] eqn:RB.
+- exists (if b then e#r1 else e#r2); split.
++ simpl. destruct b; auto.
++ destruct (eval_condition c e##args m) as [b'|] eqn:EC; simpl; auto.
+ assert (b = b').
+ { eapply resolve_branch_sound; eauto.
+ rewrite <- EC. apply eval_static_condition_sound with bc.
+ subst vl. exact (aregs_sound _ _ _ args MATCH). }
+ subst b'. apply Val.lessdef_normalize.
+- generalize (cond_strength_reduction_correct c args vl H).
+ destruct (cond_strength_reduction c args vl) as [cond' args']; intros EQ.
+ econstructor; split. simpl; eauto. rewrite EQ; auto.
+Qed.
+
Lemma make_addimm_correct:
forall n r,
let (op, args) := make_addimm n r in
@@ -905,6 +927,8 @@ Proof.
auto.
(* cond *)
inv H0. apply make_cmp_correct; auto.
+(* select *)
+ inv H0. apply make_select_correct; congruence.
(* mulf *)
InvApproxRegs; SimplVM; inv H0. rewrite <- H2. apply make_mulfimm_correct; auto.
InvApproxRegs; SimplVM; inv H0. fold (Val.mulf (Vfloat n1) e#r2).
diff --git a/x86/SelectOp.vp b/x86/SelectOp.vp
index c0434a67..a23c37d5 100644
--- a/x86/SelectOp.vp
+++ b/x86/SelectOp.vp
@@ -38,9 +38,10 @@
Require Import Coqlib.
Require Import Compopts.
-Require Import AST Integers Floats.
+Require Import AST Integers Floats Builtins.
Require Import Op CminorSel.
Require Import OpHelpers.
+Require Archi.
Local Open Scope cminorsel_scope.
@@ -499,21 +500,27 @@ Nondetfunction floatofint (e: expr) :=
end.
Definition intuoffloat (e: expr) :=
- Elet e
- (Elet (Eop (Ofloatconst (Float.of_intu Float.ox8000_0000)) Enil)
- (Econdition (CEcond (Ccompf Clt) (Eletvar 1 ::: Eletvar 0 ::: Enil))
- (intoffloat (Eletvar 1))
- (addimm Float.ox8000_0000 (intoffloat (subf (Eletvar 1) (Eletvar 0))))))%nat.
+ if Archi.splitlong then
+ Elet e
+ (Elet (Eop (Ofloatconst (Float.of_intu Float.ox8000_0000)) Enil)
+ (Econdition (CEcond (Ccompf Clt) (Eletvar 1 ::: Eletvar 0 ::: Enil))
+ (intoffloat (Eletvar 1))
+ (addimm Float.ox8000_0000 (intoffloat (subf (Eletvar 1) (Eletvar 0))))))%nat
+ else
+ Eop Olowlong (Eop Olongoffloat (e ::: Enil) ::: Enil).
Nondetfunction floatofintu (e: expr) :=
match e with
| Eop (Ointconst n) Enil => Eop (Ofloatconst (Float.of_intu n)) Enil
| _ =>
- let f := Eop (Ofloatconst (Float.of_intu Float.ox8000_0000)) Enil in
- Elet e
- (Econdition (CEcond (Ccompuimm Clt Float.ox8000_0000) (Eletvar O ::: Enil))
- (floatofint (Eletvar O))
- (addf (floatofint (addimm (Int.neg Float.ox8000_0000) (Eletvar O))) f))
+ if Archi.splitlong then
+ let f := Eop (Ofloatconst (Float.of_intu Float.ox8000_0000)) Enil in
+ Elet e
+ (Econdition (CEcond (Ccompuimm Clt Float.ox8000_0000) (Eletvar O ::: Enil))
+ (floatofint (Eletvar O))
+ (addf (floatofint (addimm (Int.neg Float.ox8000_0000) (Eletvar O))) f))
+ else
+ Eop Ofloatoflong (Eop Ocast32unsigned (e ::: Enil) ::: Enil)
end.
Definition intofsingle (e: expr) := Eop Ointofsingle (e ::: Enil).
@@ -576,3 +583,8 @@ Definition divf_base (e1: expr) (e2: expr) :=
Definition divfs_base (e1: expr) (e2: expr) :=
Eop Odivfs (e1 ::: e2 ::: Enil).
+
+(** Platform-specific known builtins *)
+
+Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr :=
+ None.
diff --git a/x86/SelectOpproof.v b/x86/SelectOpproof.v
index 5e0f84e3..5703c758 100644
--- a/x86/SelectOpproof.v
+++ b/x86/SelectOpproof.v
@@ -13,15 +13,9 @@
(** Correctness of instruction selection for operators *)
Require Import Coqlib.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import Globalenvs.
-Require Import Cminor.
-Require Import Op.
-Require Import CminorSel.
+Require Import AST Integers Floats.
+Require Import Values Memory Builtins Globalenvs.
+Require Import Cminor Op CminorSel.
Require Import SelectOp.
Require Import OpHelpers.
Require Import OpHelpersproof.
@@ -838,7 +832,8 @@ Proof.
intros. destruct x; simpl in H0; try discriminate.
destruct (Float.to_intu f) as [n|] eqn:?; simpl in H0; inv H0.
exists (Vint n); split; auto. unfold intuoffloat.
- set (im := Int.repr Int.half_modulus).
+ destruct Archi.splitlong.
+- set (im := Int.repr Int.half_modulus).
set (fm := Float.of_intu im).
assert (eval_expr ge sp e m (Vfloat fm :: Vfloat f :: le) (Eletvar (S O)) (Vfloat f)).
constructor. auto.
@@ -865,6 +860,11 @@ Proof.
rewrite Int.add_neg_zero in A4.
rewrite Int.add_zero in A4.
auto.
+- apply Float.to_intu_to_long in Heqo. repeat econstructor. eauto.
+ simpl. rewrite Heqo; reflexivity.
+ simpl. unfold Int64.loword. rewrite Int64.unsigned_repr, Int.repr_unsigned; auto.
+ assert (Int.modulus < Int64.max_unsigned) by reflexivity.
+ generalize (Int.unsigned_range n); omega.
Qed.
Theorem eval_floatofintu:
@@ -874,10 +874,11 @@ Theorem eval_floatofintu:
exists v, eval_expr ge sp e m le (floatofintu a) v /\ Val.lessdef y v.
Proof.
intros until y; unfold floatofintu. case (floatofintu_match a); intros.
- InvEval. TrivialExists.
- destruct x; simpl in H0; try discriminate. inv H0.
+- InvEval. TrivialExists.
+- destruct x; simpl in H0; try discriminate. inv H0.
exists (Vfloat (Float.of_intu i)); split; auto.
- econstructor. eauto.
+ destruct Archi.splitlong.
++ econstructor. eauto.
set (fm := Float.of_intu Float.ox8000_0000).
assert (eval_expr ge sp e m (Vint i :: le) (Eletvar O) (Vint i)).
constructor. auto.
@@ -893,6 +894,7 @@ Proof.
constructor. EvalOp. simpl; eauto. constructor. simpl; eauto.
fold fm. rewrite Float.of_intu_of_int_2; auto.
rewrite Int.sub_add_opp. auto.
++ rewrite Float.of_intu_of_long. repeat econstructor. eauto. reflexivity.
Qed.
Theorem eval_intofsingle:
@@ -1014,7 +1016,6 @@ Proof.
- constructor; auto.
Qed.
-
(* floating-point division without HELPERS *)
Theorem eval_divf_base:
forall le a b x y,
@@ -1035,4 +1036,17 @@ Proof.
intros; unfold divfs_base.
TrivialExists.
Qed.
+
+(** Platform-specific known builtins *)
+
+Theorem eval_platform_builtin:
+ forall bf al a vl v le,
+ platform_builtin bf al = Some a ->
+ eval_exprlist ge sp e m le al vl ->
+ platform_builtin_sem bf vl = Some v ->
+ exists v', eval_expr ge sp e m le a v' /\ Val.lessdef v v'.
+Proof.
+ intros. discriminate.
+Qed.
+
End CMCONSTR.
diff --git a/x86/TargetPrinter.ml b/x86/TargetPrinter.ml
index 3025d2e7..6159437e 100644
--- a/x86/TargetPrinter.ml
+++ b/x86/TargetPrinter.ml
@@ -62,8 +62,8 @@ let ireg64 oc r = output_string oc (int64_reg_name r)
let ireg = if Archi.ptr64 then ireg64 else ireg32
let freg oc r = output_string oc (float_reg_name r)
-let preg oc = function
- | IR r -> ireg oc r
+let preg_asm oc ty = function
+ | IR r -> if ty = Tlong then ireg64 oc r else ireg32 oc r
| FR r -> freg oc r
| _ -> assert false
@@ -399,7 +399,13 @@ module Target(System: SYSTEM):TARGET =
(* Printing of instructions *)
-(* Reminder on AT&T syntax: op source, dest *)
+(* Reminder on X86 assembly syntaxes:
+ AT&T syntax Intel syntax
+ (used by GNU as) (used in reference manuals)
+ dst <- op(src) op src, dst op dst, src
+ dst <- op(dst, src2) op src2, dst op dst, src2
+ dst <- op(dst, src2, src3) op src3, src2, dst op dst, src2, src3
+*)
let print_instruction oc = function
(* Moves *)
@@ -752,29 +758,29 @@ module Target(System: SYSTEM):TARGET =
| Pcfi_adjust sz ->
cfi_adjust oc (camlint_of_coqint sz)
| Pfmadd132 (res,a1,a2) ->
- fprintf oc " vfmadd132sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfmadd132sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pfmadd213 (res,a1,a2) ->
- fprintf oc " vfmadd213sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfmadd213sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pfmadd231 (res,a1,a2) ->
- fprintf oc " vfmadd231sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfmadd231sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pfmsub132 (res,a1,a2) ->
- fprintf oc " vfmsub132sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfmsub132sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pfmsub213 (res,a1,a2) ->
- fprintf oc " vfmsub213sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfmsub213sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pfmsub231 (res,a1,a2) ->
- fprintf oc " vfmsub231sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfmsub231sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pfnmadd132 (res,a1,a2) ->
- fprintf oc " vfnmadd132sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfnmadd132sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pfnmadd213 (res,a1,a2) ->
- fprintf oc " vfnmadd213sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfnmadd213sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pfnmadd231 (res,a1,a2) ->
- fprintf oc " vfnmadd231sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfnmadd231sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pfnmsub132 (res,a1,a2) ->
- fprintf oc " vfnmsub132sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfnmsub132sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pfnmsub213 (res,a1,a2) ->
- fprintf oc " vfnmsub213sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfnmsub213sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pfnmsub231 (res,a1,a2) ->
- fprintf oc " vfnmsub231sd %a, %a, %a\n" freg a1 freg a2 freg res
+ fprintf oc " vfnmsub231sd %a, %a, %a\n" freg a2 freg a1 freg res
| Pmaxsd (res,a1) ->
fprintf oc " maxsd %a, %a\n" freg a1 freg res
| Pminsd (res,a1) ->
@@ -826,7 +832,7 @@ module Target(System: SYSTEM):TARGET =
(P.to_int kind) (extern_atom txt) args
| EF_inline_asm(txt, sg, clob) ->
fprintf oc "%s begin inline assembly\n\t" comment;
- print_inline_asm preg oc (camlstring_of_coqstring txt) sg args res;
+ print_inline_asm preg_asm oc (camlstring_of_coqstring txt) sg args res;
fprintf oc "%s end inline assembly\n" comment
| _ ->
assert false
diff --git a/x86_32/Archi.v b/x86_32/Archi.v
index f10570e2..e9d05c14 100644
--- a/x86_32/Archi.v
+++ b/x86_32/Archi.v
@@ -16,7 +16,7 @@
(** Architecture-dependent parameters for x86 in 32-bit mode *)
-Require Import ZArith.
+Require Import ZArith List.
(*From Flocq*)
Require Import Binary Bits.
@@ -34,24 +34,33 @@ Proof.
unfold splitlong. destruct ptr64; simpl; congruence.
Qed.
-Definition default_nan_64 : { x : binary64 | is_nan _ _ x = true } :=
- exist _ (B754_nan 53 1024 true (iter_nat 51 _ xO xH) (eq_refl true)) (eq_refl true).
+Definition default_nan_64 := (true, iter_nat 51 _ xO xH).
+Definition default_nan_32 := (true, iter_nat 22 _ xO xH).
-Definition choose_binop_pl_64 (pl1 pl2 : positive) :=
- false. (**r always choose first NaN *)
+(* Always choose the first NaN argument, if any *)
-Definition default_nan_32 : { x : binary32 | is_nan _ _ x = true } :=
- exist _ (B754_nan 24 128 true (iter_nat 22 _ xO xH) (eq_refl true)) (eq_refl true).
+Definition choose_nan_64 (l: list (bool * positive)) : bool * positive :=
+ match l with nil => default_nan_64 | n :: _ => n end.
-Definition choose_binop_pl_32 (pl1 pl2 : positive) :=
- false. (**r always choose first NaN *)
+Definition choose_nan_32 (l: list (bool * positive)) : bool * positive :=
+ match l with nil => default_nan_32 | n :: _ => n end.
-Definition fpu_returns_default_qNaN := false.
+Lemma choose_nan_64_idem: forall n,
+ choose_nan_64 (n :: n :: nil) = choose_nan_64 (n :: nil).
+Proof. auto. Qed.
+
+Lemma choose_nan_32_idem: forall n,
+ choose_nan_32 (n :: n :: nil) = choose_nan_32 (n :: nil).
+Proof. auto. Qed.
+
+Definition fma_order {A: Type} (x y z: A) := (x, y, z).
+
+Definition fma_invalid_mul_is_nan := false.
Definition float_of_single_preserves_sNaN := false.
Global Opaque ptr64 big_endian splitlong
- default_nan_64 choose_binop_pl_64
- default_nan_32 choose_binop_pl_32
- fpu_returns_default_qNaN
+ default_nan_64 choose_nan_64
+ default_nan_32 choose_nan_32
+ fma_order fma_invalid_mul_is_nan
float_of_single_preserves_sNaN.
diff --git a/x86_64/Archi.v b/x86_64/Archi.v
index 01eb36dd..959d8dc1 100644
--- a/x86_64/Archi.v
+++ b/x86_64/Archi.v
@@ -16,7 +16,7 @@
(** Architecture-dependent parameters for x86 in 64-bit mode *)
-Require Import ZArith.
+Require Import ZArith List.
(*From Flocq*)
Require Import Binary Bits.
@@ -34,24 +34,33 @@ Proof.
unfold splitlong. destruct ptr64; simpl; congruence.
Qed.
-Definition default_nan_64 : { x : binary64 | is_nan _ _ x = true } :=
- exist _ (B754_nan 53 1024 true (iter_nat 51 _ xO xH) (eq_refl true)) (eq_refl true).
+Definition default_nan_64 := (true, iter_nat 51 _ xO xH).
+Definition default_nan_32 := (true, iter_nat 22 _ xO xH).
-Definition choose_binop_pl_64 (pl1 pl2 : positive) :=
- false. (**r always choose first NaN *)
+(* Always choose the first NaN argument, if any *)
-Definition default_nan_32 : { x : binary32 | is_nan _ _ x = true } :=
- exist _ (B754_nan 24 128 true (iter_nat 22 _ xO xH) (eq_refl true)) (eq_refl true).
+Definition choose_nan_64 (l: list (bool * positive)) : bool * positive :=
+ match l with nil => default_nan_64 | n :: _ => n end.
-Definition choose_binop_pl_32 (pl1 pl2 : positive) :=
- false. (**r always choose first NaN *)
+Definition choose_nan_32 (l: list (bool * positive)) : bool * positive :=
+ match l with nil => default_nan_32 | n :: _ => n end.
-Definition fpu_returns_default_qNaN := false.
+Lemma choose_nan_64_idem: forall n,
+ choose_nan_64 (n :: n :: nil) = choose_nan_64 (n :: nil).
+Proof. auto. Qed.
+
+Lemma choose_nan_32_idem: forall n,
+ choose_nan_32 (n :: n :: nil) = choose_nan_32 (n :: nil).
+Proof. auto. Qed.
+
+Definition fma_order {A: Type} (x y z: A) := (x, y, z).
+
+Definition fma_invalid_mul_is_nan := false.
Definition float_of_single_preserves_sNaN := false.
Global Opaque ptr64 big_endian splitlong
- default_nan_64 choose_binop_pl_64
- default_nan_32 choose_binop_pl_32
- fpu_returns_default_qNaN
+ default_nan_64 choose_nan_64
+ default_nan_32 choose_nan_32
+ fma_order fma_invalid_mul_is_nan
float_of_single_preserves_sNaN.