summaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorGravatar Stephane Glondu <steph@glondu.net>2010-12-01 13:33:59 +0100
committerGravatar Stephane Glondu <steph@glondu.net>2010-12-01 13:33:59 +0100
commitaa51449cb692a9a1d183b2663679645aba16b036 (patch)
treef32b2fd38d3442e4d583009cb09cbac701c33bc0
parenta3c2b799e0eceb0896af062887c762c654d2e2f6 (diff)
parent8ab748064ddeec8400859e210bf9963826cba631 (diff)
Merge commit 'upstream/0.2.1'
-rw-r--r--.gitignore2
-rw-r--r--AAC.v1120
-rw-r--r--AAC_coq.ml591
-rw-r--r--AAC_coq.mli231
-rw-r--r--AAC_helper.ml41
-rw-r--r--AAC_helper.mli33
-rw-r--r--AAC_matcher.ml1206
-rw-r--r--AAC_matcher.mli (renamed from matcher.mli)129
-rw-r--r--AAC_print.ml100
-rw-r--r--AAC_print.mli23
-rw-r--r--AAC_rewrite.ml524
-rw-r--r--AAC_rewrite.mli9
-rw-r--r--AAC_search_monad.ml65
-rw-r--r--AAC_search_monad.mli41
-rw-r--r--AAC_theory.ml1097
-rw-r--r--AAC_theory.mli (renamed from theory.mli)188
-rw-r--r--Caveats.v377
-rw-r--r--Instances.v263
-rw-r--r--README.txt29
-rw-r--r--Tests.v373
-rw-r--r--Tutorial.v416
-rw-r--r--aac_rewrite.ml308
-rw-r--r--aac_rewrite.mli59
-rw-r--r--coq.ml212
-rw-r--r--coq.mli117
-rw-r--r--files.txt12
-rw-r--r--magic.txt1
-rwxr-xr-xmake_makefile37
-rw-r--r--matcher.ml980
-rw-r--r--theory.ml708
30 files changed, 5754 insertions, 3538 deletions
diff --git a/.gitignore b/.gitignore
index 845ca06..ce3c045 100644
--- a/.gitignore
+++ b/.gitignore
@@ -1 +1,3 @@
+doc
+html
.pc
diff --git a/AAC.v b/AAC.v
index bdbe1b0..0f5e85d 100644
--- a/AAC.v
+++ b/AAC.v
@@ -13,35 +13,35 @@
quoted) expressions (with morphisms).
We then define a reflexive decision procedure to decide the
- equality of reified terms: first normalize reified terms, then
+ equality of reified terms: first normalise reified terms, then
compare them. This allows us to close transitivity steps
automatically, in the [aac_rewrite] tactic.
-
+
We restrict ourselves to the case where all symbols operate on a
single fixed type. In particular, this means that we cannot handle
- situations like
+ situations like
[H: forall x y, nat_of_pos (pos_of_nat (x) + y) + x = ....]
- where one occurence of [+] operates on nat while the other one
+ where one occurrence of [+] operates on nat while the other one
operates on positive. *)
-Require Import Arith NArith.
+Require Import Arith NArith.
Require Import List.
Require Import FMapPositive FMapFacts.
Require Import RelationClasses Equality.
Require Export Morphisms.
Set Implicit Arguments.
-Unset Automatic Introduction.
-Open Scope signature_scope.
+
+Local Open Scope signature_scope.
(** * Environments for the reification process: we use positive maps to index elements *)
Section sigma.
Definition sigma := PositiveMap.t.
Definition sigma_get A (null : A) (map : sigma A) (n : positive) : A :=
- match PositiveMap.find n map with
+ match PositiveMap.find n map with
| None => null
| Some x => x
end.
@@ -49,123 +49,78 @@ Section sigma.
Definition sigma_empty := @PositiveMap.empty.
End sigma.
-(** * Packaging structures *)
-
-(** ** Free symbols *)
-Module Sym.
- Section t.
- Context {X} {R : relation X} .
-
- (** type of an arity *)
- Fixpoint type_of (n: nat) :=
- match n with
- | O => X
- | S n => X -> type_of n
- end.
-
- (** relation to be preserved at an arity *)
- Fixpoint rel_of n : relation (type_of n) :=
- match n with
- | O => R
- | S n => respectful R (rel_of n)
- end.
-
- (** a symbol package contains an arity,
- a value of the corresponding type,
- and a proof that the value is a proper morphism *)
- Record pack : Type := mkPack {
- ar : nat;
- value : type_of ar;
- morph : Proper (rel_of ar) value
- }.
-
- (** helper to build default values, when filling reification environments *)
- Program Definition null (H: Equivalence R) (x : X) := mkPack 0 x _.
-
- End t.
-
-End Sym.
-
-(** ** Associative commutative operations (commutative monoids)
-
- See #<a href="Instances.html">Instances.v</a># for concrete instances of these classes.
-*)
+(** * Classes for properties of operators *)
-Class Op_AC (X:Type) (R: relation X) (plus: X -> X -> X) (zero:X) := {
- plus_compat:> Proper (R ==> R ==> R) plus;
- plus_neutral_left: forall x, R (plus zero x) x;
- plus_assoc: forall x y z, R (plus x (plus y z)) (plus (plus x y) z);
- plus_com: forall x y, R (plus x y) (plus y x)
+Class Associative (X:Type) (R:relation X) (dot: X -> X -> X) :=
+ law_assoc : forall x y z, R (dot x (dot y z)) (dot (dot x y) z).
+Class Commutative (X:Type) (R: relation X) (plus: X -> X -> X) :=
+ law_comm: forall x y, R (plus x y) (plus y x).
+Class Unit (X:Type) (R:relation X) (op : X -> X -> X) (unit:X) := {
+ law_neutral_left: forall x, R (op unit x) x;
+ law_neutral_right: forall x, R (op x unit) x
}.
-Module Sum.
- Section t.
- Context {X : Type} {R: relation X}.
- Class pack : Type := mkPack {
- plus : X -> X -> X;
- zero : X;
- opac :> @Op_AC X R plus zero
- }.
- End t.
-End Sum.
-
-
-
-(** ** Associative operations (monoids) *)
-Class Op_A (X:Type) (R:relation X) (dot: X -> X -> X) (one: X) := {
- dot_compat:> Proper (R ==> R ==> R) dot;
- dot_neutral_left: forall x, R (dot one x) x;
- dot_neutral_right: forall x, R (dot x one) x;
- dot_assoc: forall x y z, R (dot x (dot y z)) (dot (dot x y) z)
-}.
+(** Class used to find the equivalence relation on which operations
+ are A or AC, starting from the relation appearing in the goal *)
-Module Prd.
- Section t.
- Context {X : Type} {R: relation X}.
- Class pack : Type := mkPack {
- dot : X -> X -> X;
- one : X;
- opa :> @Op_A X R dot one
- }.
- End t.
-End Prd.
+Class AAC_lift X (R: relation X) (E : relation X) := {
+ aac_lift_equivalence : Equivalence E;
+ aac_list_proper : Proper (E ==> E ==> iff) R
+}.
+(** simple instances, when we have a subrelation, or an equivalence *)
-(** an helper lemma to derive an A class out of an AC one (not
- declared as an instance to avoid useless branches in typeclass
- resolution) *)
-Lemma AC_A {X} {R:relation X} {EQ: Equivalence R} {plus} {zero} (op: Op_AC R plus zero): Op_A R plus zero.
+Instance aac_lift_subrelation {X} {R} {E} {HE: Equivalence E}
+ {HR: @Transitive X R} {HER: subrelation E R}: AAC_lift R E | 3.
Proof.
- split; try apply op.
- intros. rewrite plus_com. apply op.
+ constructor; trivial.
+ intros ? ? H ? ? H'. split; intro G.
+ rewrite <- H, G. apply HER, H'.
+ rewrite H, G. apply HER. symmetry. apply H'.
Qed.
+Instance aac_lift_proper {X} {R : relation X} {E} {HE: Equivalence E}
+ {HR: Proper (E==>E==>iff) R}: AAC_lift R E | 4.
+
+Module Internal.
+
(** * Utilities for the evaluation function *)
+
Section copy.
- Context {X} {R: relation X} {HR : @Equivalence X R} {plus} {zero} {op: @Op_AC X R plus zero}.
+ Context {X} {R} {HR: @Equivalence X R} {plus}
+ (op: Associative R plus) (op': Commutative R plus) (po: Proper (R ==> R ==> R) plus).
(* copy n x = x+...+x (n times) *)
- Definition copy n x := Prect (fun _ => X) x (fun _ xn => plus x xn) n.
-
+ Fixpoint copy' n x := match n with
+ | xH => x
+ | xI n => let xn := copy' n x in plus (plus xn xn) x
+ | xO n => let xn := copy' n x in (plus xn xn)
+ end.
+ Definition copy n x := Prect (fun _ => X) x (fun _ xn => plus x xn) n.
+
Lemma copy_plus : forall n m x, R (copy (n+m) x) (plus (copy n x) (copy m x)).
Proof.
unfold copy.
induction n using Pind; intros m x.
- rewrite Prect_base. rewrite <- Pplus_one_succ_l. rewrite Prect_succ. reflexivity.
- rewrite Pplus_succ_permute_l. rewrite 2Prect_succ. rewrite IHn. apply plus_assoc.
+ rewrite Prect_base. rewrite <- Pplus_one_succ_l. rewrite Prect_succ. reflexivity.
+ rewrite Pplus_succ_permute_l. rewrite 2Prect_succ. rewrite IHn. apply op.
Qed.
+ Lemma copy_xH : forall x, R (copy 1 x) x.
+ Proof. intros; unfold copy; rewrite Prect_base. reflexivity. Qed.
+ Lemma copy_Psucc : forall n x, R (copy (Psucc n) x) (plus x (copy n x)).
+ Proof. intros; unfold copy; rewrite Prect_succ. reflexivity. Qed.
Global Instance copy_compat n: Proper (R ==> R) (copy n).
Proof.
unfold copy.
induction n using Pind; intros x y H.
- rewrite 2Prect_base. assumption.
- rewrite 2Prect_succ. apply plus_compat; auto.
+ rewrite 2Prect_base. assumption.
+ rewrite 2Prect_succ. apply po; auto.
Qed.
End copy.
@@ -173,12 +128,12 @@ End copy.
(** * Utilities for positive numbers
which we use as:
- indices for morphisms and symbols
- - multiplicity of terms in sums
- *)
-Notation idx := positive.
+ - multiplicity of terms in sums *)
+
+Local Notation idx := positive.
Fixpoint eq_idx_bool i j :=
- match i,j with
+ match i,j with
| xH, xH => true
| xO i, xO j => eq_idx_bool i j
| xI i, xI j => eq_idx_bool i j
@@ -186,7 +141,7 @@ Fixpoint eq_idx_bool i j :=
end.
Fixpoint idx_compare i j :=
- match i,j with
+ match i,j with
| xH, xH => Eq
| xH, _ => Lt
| _, xH => Gt
@@ -196,9 +151,9 @@ Fixpoint idx_compare i j :=
| xO _, xI _ => Lt
end.
-Notation pos_compare := idx_compare (only parsing).
+Local Notation pos_compare := idx_compare (only parsing).
-(** specification predicate for boolean binary functions *)
+(** Specification predicate for boolean binary functions *)
Inductive decide_spec {A} {B} (R : A -> B -> Prop) (x : A) (y : B) : bool -> Prop :=
| decide_true : R x y -> decide_spec R x y true
| decide_false : ~(R x y) -> decide_spec R x y false.
@@ -208,7 +163,7 @@ Proof.
induction i; destruct j; simpl; try (constructor; congruence).
case (IHi j); constructor; congruence.
case (IHi j); constructor; congruence.
-Qed.
+Qed.
(** weak specification predicate for comparison functions: only the 'Eq' case is specified *)
Inductive compare_weak_spec A: A -> A -> comparison -> Prop :=
@@ -222,27 +177,29 @@ Proof. induction i; destruct j; simpl; try constructor; case (IHi j); intros; co
Lemma idx_compare_reflect_eq: forall i j, idx_compare i j = Eq -> i=j.
Proof. intros i j. case (pos_compare_weak_spec i j); intros; congruence. Qed.
-
(** * Dependent types utilities *)
-Notation cast T H u := (eq_rect _ T u _ H).
+
+Local Notation cast T H u := (eq_rect _ T u _ H).
Section dep.
Variable U: Type.
Variable T: U -> Type.
- Lemma cast_eq: (forall u v: U, {u=v}+{u<>v}) ->
+ Lemma cast_eq: (forall u v: U, {u=v}+{u<>v}) ->
forall A (H: A=A) (u: T A), cast T H u = u.
- Proof. intros. rewrite <- Eqdep_dec.eq_rect_eq_dec; trivial. Qed.
+ Proof. intros. rewrite <- Eqdep_dec.eq_rect_eq_dec; trivial. Qed.
Variable f: forall A B, T A -> T B -> comparison.
Definition reflect_eqdep := forall A u B v (H: A=B), @f A B u v = Eq -> cast T H u = v.
(* these lemmas have to remain transparent to get structural recursion
in the lemma [tcompare_weak_spec] below *)
- Lemma reflect_eqdep_eq: reflect_eqdep -> forall A u v, @f A A u v = Eq -> u = v.
- Proof. intros H A u v He. apply (H _ _ _ _ eq_refl He). Defined.
+ Lemma reflect_eqdep_eq: reflect_eqdep ->
+ forall A u v, @f A A u v = Eq -> u = v.
+ Proof. intros H A u v He. apply (H _ _ _ _ eq_refl He). Defined.
- Lemma reflect_eqdep_weak_spec: reflect_eqdep -> forall A u v, compare_weak_spec u v (@f A A u v).
+ Lemma reflect_eqdep_weak_spec: reflect_eqdep ->
+ forall A u v, compare_weak_spec u v (@f A A u v).
Proof.
intros. case_eq (f u v); try constructor.
intro H'. apply reflect_eqdep_eq in H'. subst. constructor. assumption.
@@ -250,42 +207,70 @@ Section dep.
End dep.
-(** * Utilities about lists and multisets *)
+
+(** * Utilities about (non-empty) lists and multisets *)
+
+Inductive nelist (A : Type) : Type :=
+| nil : A -> nelist A
+| cons : A -> nelist A -> nelist A.
+
+Local Notation "x :: y" := (cons x y).
+
+Fixpoint nelist_map (A B: Type) (f: A -> B) l :=
+ match l with
+ | nil x => nil ( f x)
+ | cons x l => cons ( f x) (nelist_map f l)
+ end.
+
+Fixpoint appne A l l' : nelist A :=
+ match l with
+ nil x => cons x l'
+ | cons t q => cons t (appne A q l')
+ end.
+
+Local Notation "x ++ y" := (appne x y).
(** finite multisets are represented with ordered lists with multiplicities *)
-Definition mset A := list (A*positive).
+Definition mset A := nelist (A*positive).
(** lexicographic composition of comparisons (this is a notation to keep it lazy) *)
-Notation lex e f := (match e with Eq => f | _ => e end).
+Local Notation lex e f := (match e with Eq => f | _ => e end).
Section lists.
- (** comparison function *)
+ (** comparison functions *)
+
Section c.
Variables A B: Type.
Variable compare: A -> B -> comparison.
- Fixpoint list_compare h k :=
+ Fixpoint list_compare h k :=
match h,k with
- | nil, nil => Eq
- | nil, _ => Lt
- | _, nil => Gt
- | u::h, v::k => lex (compare u v) (list_compare h k)
- end.
+ | nil x, nil y => compare x y
+ | nil x, _ => Lt
+ | _, nil x => Gt
+ | u::h, v::k => lex (compare u v) (list_compare h k)
+ end.
End c.
Definition mset_compare A B compare: mset A -> mset B -> comparison :=
- list_compare (fun un vm => let '(u,n) := un in let '(v,m) := vm in lex (compare u v) (pos_compare n m)).
+ list_compare (fun un vm =>
+ let '(u,n) := un in
+ let '(v,m) := vm in
+ lex (compare u v) (pos_compare n m)).
Section list_compare_weak_spec.
Variable A: Type.
Variable compare: A -> A -> comparison.
Hypothesis Hcompare: forall u v, compare_weak_spec u v (compare u v).
- (* this lemma has to remain transparent to get structural recursion
+ (* this lemma has to remain transparent to get structural recursion
in the lemma [tcompare_weak_spec] below *)
- Lemma list_compare_weak_spec: forall h k, compare_weak_spec h k (list_compare compare h k).
+ Lemma list_compare_weak_spec: forall h k,
+ compare_weak_spec h k (list_compare compare h k).
Proof.
induction h as [|u h IHh]; destruct k as [|v k]; simpl; try constructor.
- case (Hcompare u v); try constructor.
+
+ case (Hcompare a a0 ); try constructor.
+ case (Hcompare u v ); try constructor.
case (IHh k); intros; constructor.
Defined.
End list_compare_weak_spec.
@@ -294,9 +279,10 @@ Section lists.
Variable A: Type.
Variable compare: A -> A -> comparison.
Hypothesis Hcompare: forall u v, compare_weak_spec u v (compare u v).
- (* this lemma has to remain transparent to get structural recursion
+ (* this lemma has to remain transparent to get structural recursion
in the lemma [tcompare_weak_spec] below *)
- Lemma mset_compare_weak_spec: forall h k, compare_weak_spec h k (mset_compare compare h k).
+ Lemma mset_compare_weak_spec: forall h k,
+ compare_weak_spec h k (mset_compare compare h k).
Proof.
apply list_compare_weak_spec.
intros [u n] [v m].
@@ -306,17 +292,40 @@ Section lists.
End mset_compare_weak_spec.
(** (sorted) merging functions *)
+
Section m.
Variable A: Type.
Variable compare: A -> A -> comparison.
-
+ Definition insert n1 h1 :=
+ let fix insert_aux l2 :=
+ match l2 with
+ | nil (h2,n2) =>
+ match compare h1 h2 with
+ | Eq => nil (h1,Pplus n1 n2)
+ | Lt => (h1,n1):: nil (h2,n2)
+ | Gt => (h2,n2):: nil (h1,n1)
+ end
+ | (h2,n2)::t2 =>
+ match compare h1 h2 with
+ | Eq => (h1,Pplus n1 n2):: t2
+ | Lt => (h1,n1)::l2
+ | Gt => (h2,n2)::insert_aux t2
+ end
+ end
+ in insert_aux.
+
Fixpoint merge_msets l1 :=
match l1 with
- | nil => fun l2 => l2
+ | nil (h1,n1) => fun l2 => insert n1 h1 l2
| (h1,n1)::t1 =>
let fix merge_aux l2 :=
match l2 with
- | nil => l1
+ | nil (h2,n2) =>
+ match compare h1 h2 with
+ | Eq => (h1,Pplus n1 n2) :: t1
+ | Lt => (h1,n1):: merge_msets t1 l2
+ | Gt => (h2,n2):: l1
+ end
| (h2,n2)::t2 =>
match compare h1 h2 with
| Eq => (h1,Pplus n1 n2)::merge_msets t1 t2
@@ -328,44 +337,119 @@ Section lists.
end.
(** interpretation of a list with a constant and a binary operation *)
+
Variable B: Type.
Variable map: A -> B.
Variable b2: B -> B -> B.
- Variable b0: B.
Fixpoint fold_map l :=
match l with
- | nil => b0
- | u::nil => map u
+ | nil x => map x
| u::l => b2 (map u) (fold_map l)
end.
(** mapping and merging *)
- Variable merge: A -> list B -> list B.
- Fixpoint merge_map (l: list A): list B :=
+
+ Variable merge: A -> nelist B -> nelist B.
+ Fixpoint merge_map (l: nelist A): nelist B :=
match l with
- | nil => nil
+ | nil x => nil (map x)
| u::l => merge u (merge_map l)
end.
- End m.
+ Variable ret : A -> B.
+ Variable bind : A -> B -> B.
+ Fixpoint fold_map' (l : nelist A) : B :=
+ match l with
+ | nil x => ret x
+ | u::l => bind u (fold_map' l)
+ end.
+ End m.
End lists.
+(** * Packaging structures *)
+
+(** ** free symbols *)
+
+Module Sym.
+ Section t.
+ Context {X} {R : relation X} .
+
+ (** type of an arity *)
+ Fixpoint type_of (n: nat) :=
+ match n with
+ | O => X
+ | S n => X -> type_of n
+ end.
+
+ (** relation to be preserved at an arity *)
+ Fixpoint rel_of n : relation (type_of n) :=
+ match n with
+ | O => R
+ | S n => respectful R (rel_of n)
+ end.
+
+ (** a symbol package contains an arity,
+ a value of the corresponding type,
+ and a proof that the value is a proper morphism *)
+ Record pack : Type := mkPack {
+ ar : nat;
+ value :> type_of ar;
+ morph : Proper (rel_of ar) value
+ }.
+
+ (** helper to build default values, when filling reification environments *)
+ Definition null: pack := mkPack 1 (fun x => x) (fun _ _ H => H).
+
+ End t.
+
+End Sym.
+
+(** ** binary operations *)
+
+Module Bin.
+ Section t.
+ Context {X} {R: relation X}.
+
+ Record pack := mk_pack {
+ value:> X -> X -> X;
+ compat: Proper (R ==> R ==> R) value;
+ assoc: Associative R value;
+ comm: option (Commutative R value)
+ }.
+ End t.
+ (* See #<a href="Instances.html">Instances.v</a># for concrete instances of these classes. *)
+
+End Bin.
(** * Reification, normalisation, and decision *)
+
Section s.
Context {X} {R: relation X} {E: @Equivalence X R}.
Infix "==" := R (at level 80).
(* We use environments to store the various operators and the
morphisms.*)
-
+
Variable e_sym: idx -> @Sym.pack X R.
- Variable e_prd: idx -> @Prd.pack X R.
- Variable e_sum: idx -> @Sum.pack X R.
- Hint Resolve e_sum e_prd: typeclass_instances.
+ Variable e_bin: idx -> @Bin.pack X R.
+
+ (** packaging units (depends on e_bin) *)
+
+ Record unit_of u := mk_unit_for {
+ uf_idx: idx;
+ uf_desc: Unit R (Bin.value (e_bin uf_idx)) u
+ }.
+
+ Record unit_pack := mk_unit_pack {
+ u_value:> X;
+ u_desc: list (unit_of u_value)
+ }.
+ Variable e_unit: positive -> unit_pack.
+
+ Hint Resolve e_bin e_unit: typeclass_instances.
(** ** Almost normalised syntax
a term in [T] is in normal form if:
@@ -373,16 +457,18 @@ Section s.
- products do not contain products
- there are no unary sums or products
- lists and msets are lexicographically sorted according to the order we define below
-
- [vT n] denotes the set of term vectors of size [n] (the mutual dependency could be removed),
+
+ [vT n] denotes the set of term vectors of size [n] (the mutual dependency could be removed),
Note that [T] and [vT] depend on the [e_sym] environment (which
contains, among other things, the arity of symbols)
*)
- Inductive T: Type :=
+
+ Inductive T: Type :=
| sum: idx -> mset T -> T
- | prd: idx -> list T -> T
+ | prd: idx -> nelist T -> T
| sym: forall i, vT (Sym.ar (e_sym i)) -> T
+ | unit : idx -> T
with vT: nat -> Type :=
| vnil: vT O
| vcons: forall n, T -> vT n -> vT (S n).
@@ -390,14 +476,18 @@ Section s.
(** lexicographic rpo over the normalised syntax *)
Fixpoint compare (u v: T) :=
- match u,v with
+ match u,v with
| sum i l, sum j vs => lex (idx_compare i j) (mset_compare compare l vs)
| prd i l, prd j vs => lex (idx_compare i j) (list_compare compare l vs)
| sym i l, sym j vs => lex (idx_compare i j) (vcompare l vs)
+ | unit i , unit j => idx_compare i j
+ | unit _ , _ => Lt
+ | _ , unit _ => Gt
| sum _ _, _ => Lt
| _ , sum _ _ => Gt
| prd _ _, _ => Lt
| _ , prd _ _ => Gt
+
end
with vcompare i j (us: vT i) (vs: vT j) :=
match us,vs with
@@ -406,12 +496,27 @@ Section s.
| _, vnil => Gt
| vcons _ u us, vcons _ v vs => lex (compare u v) (vcompare us vs)
end.
+
- (** we need to show that compare reflects equality
- (this is because we work with msets rather than lists with arities)
- *)
+ (** ** Evaluation from syntax to the abstract domain *)
+
+ Fixpoint eval u: X :=
+ match u with
+ | sum i l => let o := Bin.value (e_bin i) in
+ fold_map (fun un => let '(u,n):=un in @copy _ o n (eval u)) o l
+ | prd i l => fold_map eval (Bin.value (e_bin i)) l
+ | sym i v => eval_aux v (Sym.value (e_sym i))
+ | unit i => e_unit i
+ end
+ with eval_aux i (v: vT i): Sym.type_of i -> X :=
+ match v with
+ | vnil => fun f => f
+ | vcons _ u v => fun f => eval_aux v (f (eval u))
+ end.
+ (** we need to show that compare reflects equality (this is because
+ we work with msets rather than lists with arities) *)
Lemma tcompare_weak_spec: forall (u v : T), compare_weak_spec u v (compare u v)
with vcompare_reflect_eqdep: forall i us j vs (H: i=j), vcompare us vs = Eq -> cast vT H us = vs.
Proof.
@@ -421,269 +526,610 @@ Section s.
case (mset_compare_weak_spec compare tcompare_weak_spec m m0); intros; try constructor.
destruct v; simpl; try constructor.
case (pos_compare_weak_spec p p0); intros; try constructor.
- case (list_compare_weak_spec compare tcompare_weak_spec l l0); intros; try constructor.
+ case (list_compare_weak_spec compare tcompare_weak_spec n n0); intros; try constructor.
destruct v0; simpl; try constructor.
case_eq (idx_compare i i0); intro Hi; try constructor.
apply idx_compare_reflect_eq in Hi. symmetry in Hi. subst. (* the [symmetry] is required ! *)
case_eq (vcompare v v0); intro Hv; try constructor.
- rewrite <- (vcompare_reflect_eqdep _ _ _ _ eq_refl Hv). constructor.
+ rewrite <- (vcompare_reflect_eqdep _ _ _ _ eq_refl Hv). constructor.
+ destruct v; simpl; try constructor.
+ case_eq (idx_compare p p0); intro Hi; try constructor.
+ apply idx_compare_reflect_eq in Hi. symmetry in Hi. subst. constructor.
induction us; destruct vs; simpl; intros H Huv; try discriminate.
- apply cast_eq, eq_nat_dec.
+ apply cast_eq, eq_nat_dec.
injection H; intro Hn.
- revert Huv; case (tcompare_weak_spec t t0); intros; try discriminate.
+ revert Huv; case (tcompare_weak_spec t t0); intros; try discriminate.
symmetry in Hn. subst. (* symmetry required *)
- rewrite <- (IHus _ _ eq_refl Huv).
+ rewrite <- (IHus _ _ eq_refl Huv).
apply cast_eq, eq_nat_dec.
Qed.
- (** ** Evaluation from syntax to the abstract domain *)
- Fixpoint eval u: X :=
- match u with
- | sum i l => fold_map (fun un => let '(u,n):=un in @copy _ (@Sum.plus _ _ (e_sum i)) n (eval u))
- (@Sum.plus _ _ (e_sum i)) (@Sum.zero _ _ (e_sum i)) l
- | prd i l => fold_map eval
- (@Prd.dot _ _ (e_prd i)) (@Prd.one _ _ (e_prd i)) l
- | sym i v => @eval_aux _ v (Sym.value (e_sym i))
- end
- with eval_aux i (v: vT i): Sym.type_of i -> X :=
- match v with
- | vnil => fun f => f
- | vcons _ u v => fun f => eval_aux v (f (eval u))
- end.
-
-
Instance eval_aux_compat i (l: vT i): Proper (@Sym.rel_of X R i ==> R) (eval_aux l).
Proof.
induction l; simpl; repeat intro.
assumption.
- apply IHl, H. reflexivity.
+ apply IHl, H. reflexivity.
Qed.
+
+ (* is [i] a unit for [j] ? *)
+ Definition is_unit_of j i :=
+ List.existsb (fun p => eq_idx_bool j (uf_idx p)) (u_desc (e_unit i)).
+
+ (* is [i] commutative ? *)
+ Definition is_commutative i :=
+ match Bin.comm (e_bin i) with Some _ => true | None => false end.
+
(** ** Normalisation *)
- (** auxiliary functions, to clean up sums and products *)
- Definition sum' i (u: mset T): T :=
- match u with
- | (u,xH)::nil => u
- | _ => sum i u
- end.
- Definition is_sum i (u: T): option (mset T) :=
- match u with
- | sum j l => if eq_idx_bool j i then Some l else None
- | u => None
+ Inductive discr {A} : Type :=
+ | Is_op : A -> discr
+ | Is_unit : idx -> discr
+ | Is_nothing : discr .
+
+ (* This is called sum in the std lib *)
+ Inductive m {A} {B} :=
+ | left : A -> m
+ | right : B -> m.
+
+ Definition comp A B (merge : B -> B -> B) (l : B) (l' : @m A B) : @m A B :=
+ match l' with
+ | left _ => right l
+ | right l' => right (merge l l')
end.
- Definition copy_mset n (l: mset T): mset T :=
- match n with
- | xH => l
- | _ => List.map (fun vm => let '(v,m):=vm in (v,Pmult n m)) l
+
+ (** auxiliary functions, to clean up sums *)
+
+ Section sums.
+ Variable i : idx.
+ Variable is_unit : idx -> bool.
+
+ Definition sum' (u: mset T): T :=
+ match u with
+ | nil (u,xH) => u
+ | _ => sum i u
+ end.
+
+ Definition is_sum (u: T) : @discr (mset T) :=
+ match u with
+ | sum j l => if eq_idx_bool j i then Is_op l else Is_nothing
+ | unit j => if is_unit j then Is_unit j else Is_nothing
+ | u => Is_nothing
end.
- Definition sum_to_mset i (u: T) n :=
- match is_sum i u with
- | Some l' => copy_mset n l'
- | None => (u,n)::nil
+
+ Definition copy_mset n (l: mset T): mset T :=
+ match n with
+ | xH => l
+ | _ => nelist_map (fun vm => let '(v,m):=vm in (v,Pmult n m)) l
+ end.
+
+ Definition return_sum u n :=
+ match is_sum u with
+ | Is_nothing => right (nil (u,n))
+ | Is_op l' => right (copy_mset n l')
+ | Is_unit j => left j
+ end.
+
+ Definition add_to_sum u n (l : @m idx (mset T)) :=
+ match is_sum u with
+ | Is_nothing => comp (merge_msets compare) (nil (u,n)) l
+ | Is_op l' => comp (merge_msets compare) (copy_mset n l') l
+ | Is_unit _ => l
end.
- Definition prd' i (u: list T): T :=
- match u with
- | u::nil => u
+
+ Definition norm_msets_ norm (l: mset T) :=
+ fold_map'
+ (fun un => let '(u,n) := un in return_sum (norm u) n)
+ (fun un l => let '(u,n) := un in add_to_sum (norm u) n l) l.
+
+
+ End sums.
+
+ (** similar functions for products *)
+
+ Section prds.
+
+ Variable i : idx.
+ Variable is_unit : idx -> bool.
+ Definition prd' (u: nelist T): T :=
+ match u with
+ | nil u => u
| _ => prd i u
end.
- Definition is_prd i (u: T): option (list T) :=
- match u with
- | prd j l => if eq_idx_bool j i then Some l else None
- | u => None
- end.
- Definition prd_to_list i (u: T) :=
- match is_prd i u with
- | Some l' => l'
- | None => u::nil
- end.
+ Definition is_prd (u: T) : @discr (nelist T) :=
+ match u with
+ | prd j l => if eq_idx_bool j i then Is_op l else Is_nothing
+ | unit j => if is_unit j then Is_unit j else Is_nothing
+ | u => Is_nothing
+ end.
+
+
+ Definition return_prd u :=
+ match is_prd u with
+ | Is_nothing => right (nil (u))
+ | Is_op l' => right (l')
+ | Is_unit j => left j
+ end.
+
+ Definition add_to_prd u (l : @m idx (nelist T)) :=
+ match is_prd u with
+ | Is_nothing => comp (@appne T) (nil (u)) l
+ | Is_op l' => comp (@appne T) (l') l
+ | Is_unit _ => l
+ end.
- (** we deduce the normalisation function *)
- Fixpoint norm u :=
- match u with
- | sum i l => sum' i (merge_map (fun un l => let '(u,n):=un in
- merge_msets compare (sum_to_mset i (norm u) n) l) l)
- | prd i l => prd' i (merge_map (fun u l => prd_to_list i (norm u) ++ l) l)
+ Definition norm_lists_ norm (l : nelist T) :=
+ fold_map'
+ (fun u => return_prd (norm u))
+ (fun u l => add_to_prd (norm u) l) l.
+
+
+ End prds.
+
+
+ Definition run_list x := match x with
+ | left n => nil (unit n)
+ | right l => l
+ end.
+
+ Definition norm_lists norm i l :=
+ let is_unit := is_unit_of i in
+ run_list (norm_lists_ i is_unit norm l).
+
+ Definition run_msets x := match x with
+ | left n => nil (unit n, xH)
+ | right l => l
+ end.
+
+ Definition norm_msets norm i l :=
+ let is_unit := is_unit_of i in
+ run_msets (norm_msets_ i is_unit norm l).
+
+ Fixpoint norm u {struct u}:=
+ match u with
+ | sum i l => if is_commutative i then sum' i (norm_msets norm i l) else u
+ | prd i l => prd' i (norm_lists norm i l)
| sym i l => sym i (vnorm l)
+ | unit i => unit i
end
- with vnorm i (l: vT i): vT i :=
+ with vnorm i (l: vT i): vT i :=
match l with
| vnil => vnil
| vcons _ u l => vcons (norm u) (vnorm l)
end.
-
(** ** Correctness *)
- (** auxiliary lemmas about sums *)
- Lemma sum'_sum i (l: mset T): eval (sum' i l) ==eval (sum i l) .
- Proof.
- intros i [|? [|? ?]].
- reflexivity.
- destruct p; destruct p; simpl; reflexivity.
- destruct p; destruct p; simpl; reflexivity.
- Qed.
-
- Lemma eval_sum_cons i n a (l: mset T):
- (eval (sum i ((a,n)::l))) == (@Sum.plus _ _ (e_sum i) (@copy _ (@Sum.plus _ _ (e_sum i)) n (eval a)) (eval (sum i l))).
+ Lemma is_unit_of_Unit : forall i j : idx,
+ is_unit_of i j = true -> Unit R (Bin.value (e_bin i)) (eval (unit j)).
Proof.
- intros i n a [|[b m] l]; simpl.
- rewrite plus_com, plus_neutral_left. reflexivity.
- reflexivity.
- Qed.
-
- Lemma eval_sum_app i (h k: mset T): eval (sum i (h++k)) == @Sum.plus _ _ (e_sum i) (eval (sum i h)) (eval (sum i k)).
- Proof.
- induction h; intro k.
- simpl. rewrite plus_neutral_left. reflexivity.
- simpl app. destruct a. rewrite 2 eval_sum_cons, IHh, plus_assoc. reflexivity.
- Qed.
-
- Lemma eval_merge_sum i (h k: mset T):
- eval (sum i (merge_msets compare h k)) == @Sum.plus _ _ (e_sum i) (eval (sum i h)) (eval (sum i k)).
- Proof.
- induction h as [|[a n] h IHh]; intro k.
- simpl. rewrite plus_neutral_left. reflexivity.
- induction k as [|[b m] k IHk].
- simpl. setoid_rewrite plus_com. rewrite plus_neutral_left. reflexivity.
- simpl merge_msets. rewrite eval_sum_cons. destruct (tcompare_weak_spec a b) as [a|a b|a b].
- rewrite 2eval_sum_cons. rewrite IHh. rewrite <- plus_assoc. setoid_rewrite plus_assoc at 3.
- setoid_rewrite plus_com at 5. rewrite !plus_assoc. rewrite copy_plus. reflexivity.
-
- rewrite eval_sum_cons. rewrite IHh. apply plus_assoc.
-
- rewrite <- eval_sum_cons. setoid_rewrite eval_sum_cons at 3.
- rewrite plus_com, <- plus_assoc. rewrite plus_com in IHk. rewrite <- IHk.
- rewrite eval_sum_cons. apply plus_compat; reflexivity.
+ intros. unfold is_unit_of in H.
+ rewrite existsb_exists in H.
+ destruct H as [x [H H']].
+ revert H' ; case (eq_idx_spec); [intros H' _ ; subst| intros _ H'; discriminate].
+ simpl. destruct x. simpl. auto.
Qed.
-
- Inductive sum_spec: T -> option (mset T) -> Prop :=
- | ss_sum1: forall i l, sum_spec (sum i l) (Some l)
- | ss_sum2: forall j i l, i<>j -> sum_spec (sum i l) None
- | ss_prd: forall i l, sum_spec (prd i l) None
- | ss_sym: forall i l, sum_spec (@sym i l) None.
- Lemma is_sum_spec: forall i (a: T), sum_spec a (is_sum i a).
+
+ Instance Binvalue_Commutative i (H : is_commutative i = true) : Commutative R (@Bin.value _ _ (e_bin i) ).
Proof.
- intros i [j l|j l|j l]; simpl; try constructor.
- case eq_idx_spec; intro.
- subst. constructor.
- econstructor. eassumption.
+ unfold is_commutative in H.
+ destruct (Bin.comm (e_bin i)); auto.
+ discriminate.
Qed.
- Lemma eval_is_sum: forall i (a: T) k, is_sum i a = Some k -> eval (sum i k) = eval a.
+ Instance Binvalue_Associative i :Associative R (@Bin.value _ _ (e_bin i) ).
Proof.
- intros i [j l|j l|j l] k; simpl is_sum; try (intro; discriminate).
- case eq_idx_spec; intros ? H.
- injection H. clear H. intros <-. congruence.
- discriminate.
+ destruct ((e_bin i)); auto.
Qed.
-
- Lemma copy_mset' n (l: mset T):
- copy_mset n l = List.map (fun vm => let '(v,m):=vm in (v,Pmult n m)) l.
+
+ Instance Binvalue_Proper i : Proper (R ==> R ==> R) (@Bin.value _ _ (e_bin i) ).
Proof.
- intros. destruct n; try reflexivity.
- simpl. induction l as [|[a l] IHl]; simpl; congruence.
+ destruct ((e_bin i)); auto.
Qed.
+ Hint Resolve Binvalue_Proper Binvalue_Associative Binvalue_Commutative.
+
+ (** auxiliary lemmas about sums *)
+ Hint Resolve is_unit_of_Unit.
+ Section sum_correctness.
+ Variable i : idx.
+ Variable is_unit : idx -> bool.
+ Hypothesis is_unit_sum_Unit : forall j, is_unit j = true-> @Unit X R (Bin.value (e_bin i)) (eval (unit j)).
+
+ Inductive is_sum_spec_ind : T -> @discr (mset T) -> Prop :=
+ | is_sum_spec_op : forall j l, j = i -> is_sum_spec_ind (sum j l) (Is_op l)
+ | is_sum_spec_unit : forall j, is_unit j = true -> is_sum_spec_ind (unit j) (Is_unit j)
+ | is_sum_spec_nothing : forall u, is_sum_spec_ind u (Is_nothing).
+
+ Lemma is_sum_spec u : is_sum_spec_ind u (is_sum i is_unit u).
+ Proof.
+ unfold is_sum; case u; intros; try constructor.
+ case_eq (eq_idx_bool p i); intros; subst; try constructor; auto.
+ revert H. case eq_idx_spec; try discriminate. auto.
+ case_eq (is_unit p); intros; try constructor. auto.
+ Qed.
+
+ Instance assoc : @Associative X R (Bin.value (e_bin i)).
+ Proof.
+ destruct (e_bin i). simpl. assumption.
+ Qed.
+ Instance proper : Proper (R ==> R ==> R)(Bin.value (e_bin i)).
+ Proof.
+ destruct (e_bin i). simpl. assumption.
+ Qed.
+ Hypothesis comm : @Commutative X R (Bin.value (e_bin i)).
+
+ Lemma sum'_sum : forall (l: mset T), eval (sum' i l) ==eval (sum i l) .
+ Proof.
+ intros [[a n] | [a n] l]; destruct n; simpl; reflexivity.
+ Qed.
+
+ Lemma eval_sum_nil x:
+ eval (sum i (nil (x,xH))) == (eval x).
+ Proof. rewrite <- sum'_sum. reflexivity. Qed.
+
+ Lemma eval_sum_cons : forall n a (l: mset T),
+ (eval (sum i ((a,n)::l))) == (@Bin.value _ _ (e_bin i) (@copy _ (@Bin.value _ _ (e_bin i)) n (eval a)) (eval (sum i l))).
+ Proof.
+ intros n a [[? ? ]|[b m] l]; simpl; reflexivity.
+ Qed.
+
+ Inductive compat_sum_unit : @m idx (mset T) -> Prop :=
+ | csu_left : forall x, is_unit x = true-> compat_sum_unit (left x)
+ | csu_right : forall m, compat_sum_unit (right m)
+ .
+
+ Lemma compat_sum_unit_return x n : compat_sum_unit (return_sum i is_unit x n).
+ Proof.
+ unfold return_sum.
+ case is_sum_spec; intros; try constructor; auto.
+ Qed.
+
+ Lemma compat_sum_unit_add : forall x n h,
+ compat_sum_unit h
+ ->
+ compat_sum_unit
+ (add_to_sum i (is_unit_of i) x n
+ h).
+ Proof.
+ unfold add_to_sum;intros; inversion H;
+ case_eq (is_sum i (is_unit_of i) x);
+ intros; simpl; try constructor || eauto. apply H0.
+ Qed.
+
+ (* Hint Resolve copy_plus. : this lags because of the inference of the implicit arguments *)
+ Hint Extern 5 (copy (?n + ?m) (eval ?a) == Bin.value (copy ?n (eval ?a)) (copy ?m (eval ?a))) => apply copy_plus.
+ Hint Extern 5 (?x == ?x) => reflexivity.
+ Hint Extern 5 ( Bin.value ?x ?y == Bin.value ?y ?x) => apply Bin.comm.
+
+ Lemma eval_merge_bin : forall (h k: mset T),
+ eval (sum i (merge_msets compare h k)) == @Bin.value _ _ (e_bin i) (eval (sum i h)) (eval (sum i k)).
+ Proof.
+ induction h as [[a n]|[a n] h IHh]; intro k.
+ simpl.
+ induction k as [[b m]|[b m] k IHk]; simpl.
+ destruct (tcompare_weak_spec a b) as [a|a b|a b]; simpl; auto. apply copy_plus; auto.
+
+ destruct (tcompare_weak_spec a b) as [a|a b|a b]; simpl; auto.
+ rewrite copy_plus,law_assoc; auto.
+ rewrite IHk; clear IHk. rewrite 2 law_assoc. apply proper. apply law_comm. reflexivity.
+
+ induction k as [[b m]|[b m] k IHk]; simpl; simpl in IHh.
+ destruct (tcompare_weak_spec a b) as [a|a b|a b]; simpl.
+ rewrite (law_comm _ (copy m (eval a))), law_assoc, <- copy_plus, Pplus_comm; auto.
+ rewrite <- law_assoc, IHh. reflexivity.
+ rewrite law_comm. reflexivity.
+
+ simpl in IHk.
+ destruct (tcompare_weak_spec a b) as [a|a b|a b]; simpl.
+ rewrite IHh; clear IHh. rewrite 2 law_assoc. rewrite (law_comm _ (copy m (eval a))). rewrite law_assoc, <- copy_plus, Pplus_comm; auto.
+ rewrite IHh; clear IHh. simpl. rewrite law_assoc. reflexivity.
+ rewrite 2 (law_comm (copy m (eval b))). rewrite law_assoc. apply proper; [ | reflexivity].
+ rewrite <- IHk. reflexivity.
+ Qed.
+
+
+ Lemma copy_mset' n (l: mset T): copy_mset n l = nelist_map (fun vm => let '(v,m):=vm in (v,Pmult n m)) l.
+ Proof.
+ unfold copy_mset. destruct n; try reflexivity.
+
+ simpl. induction l as [|[a l] IHl]; simpl; try congruence. destruct a. reflexivity.
+ Qed.
+
+
+ Lemma copy_mset_succ n (l: mset T): eval (sum i (copy_mset (Psucc n) l)) == @Bin.value _ _ (e_bin i) (eval (sum i l)) (eval (sum i (copy_mset n l))).
+ Proof.
+ rewrite 2 copy_mset'.
+
+ induction l as [[a m]|[a m] l IHl].
+ simpl eval. rewrite <- copy_plus; auto. rewrite Pmult_Sn_m. reflexivity.
+
+ simpl nelist_map. rewrite ! eval_sum_cons. rewrite IHl. clear IHl.
+ rewrite Pmult_Sn_m. rewrite copy_plus; auto. rewrite <- !law_assoc.
+ apply Binvalue_Proper; try reflexivity.
+ rewrite law_comm . rewrite <- !law_assoc. apply proper; try reflexivity.
+ apply law_comm.
+ Qed.
+
+ Lemma copy_mset_copy : forall n (m : mset T), eval (sum i (copy_mset n m)) == @copy _ (@Bin.value _ _ (e_bin i)) n (eval (sum i m)).
+ Proof.
+ induction n using Pind; intros.
+
+ unfold copy_mset. rewrite copy_xH. reflexivity.
+ rewrite copy_mset_succ. rewrite copy_Psucc. rewrite IHn. reflexivity.
+ Qed.
+
+ Instance compat_sum_unit_Unit : forall p, compat_sum_unit (left p) ->
+ @Unit X R (Bin.value (e_bin i)) (eval (unit p)).
+ Proof.
+ intros.
+ inversion H. subst. auto.
+ Qed.
- Lemma copy_mset_succ i n (l: mset T):
- eval (sum i (copy_mset (Psucc n) l)) == @Sum.plus _ _ (e_sum i) (eval (sum i l)) (eval (sum i (copy_mset n l))).
- Proof.
- intros. rewrite 2copy_mset'.
- induction l as [|[a m] l IHl].
- simpl eval at 2. rewrite plus_neutral_left. destruct n; reflexivity.
- simpl map. rewrite ! eval_sum_cons. rewrite IHl.
- rewrite Pmult_Sn_m. rewrite copy_plus. rewrite <- !plus_assoc. apply plus_compat; try reflexivity.
- setoid_rewrite plus_com at 3. rewrite <- !plus_assoc. apply plus_compat; try reflexivity. apply plus_com.
- Qed.
+ Lemma copy_n_unit : forall j n, is_unit j = true ->
+ eval (unit j) == @copy _ (Bin.value (e_bin i)) n (eval (unit j)).
+ Proof.
+ intros.
+ induction n using Prect.
+ rewrite copy_xH. reflexivity.
+ rewrite copy_Psucc. rewrite <- IHn. apply is_unit_sum_Unit in H. rewrite law_neutral_left. reflexivity.
+ Qed.
+
+
+ Lemma z0 l r (H : compat_sum_unit r):
+ eval (sum i (run_msets (comp (merge_msets compare) l r))) == eval (sum i ((merge_msets compare) (l) (run_msets r))).
+ Proof.
+ unfold comp. unfold run_msets.
+ case_eq r; intros; subst.
+ rewrite eval_merge_bin; auto.
+ rewrite eval_sum_nil.
+ apply compat_sum_unit_Unit in H. rewrite law_neutral_right. reflexivity.
+ reflexivity.
+ Qed.
- Lemma eval_sum_to_mset: forall i n a, eval (sum i (sum_to_mset i a n)) == @copy _ (@Sum.plus _ _ (e_sum i)) n (eval a).
+ Lemma z1 : forall n x,
+ eval (sum i (run_msets (return_sum i (is_unit) x n )))
+ == @copy _ (@Bin.value _ _ (e_bin i)) n (eval x).
+ Proof.
+ intros. unfold return_sum. unfold run_msets.
+ case (is_sum_spec); intros; subst.
+ rewrite copy_mset_copy.
+ reflexivity.
+
+ rewrite eval_sum_nil. apply copy_n_unit. auto.
+ reflexivity.
+ Qed.
+ Lemma z2 : forall u n x, compat_sum_unit x ->
+ eval (sum i ( run_msets
+ (add_to_sum i (is_unit) u n x)))
+ ==
+ @Bin.value _ _ (e_bin i) (@copy _ (@Bin.value _ _ (e_bin i)) n (eval u)) (eval (sum i (run_msets x))).
+ Proof.
+ intros u n x Hix .
+ unfold add_to_sum.
+ case is_sum_spec; intros; subst.
+
+ rewrite z0 by auto. rewrite eval_merge_bin. rewrite copy_mset_copy. reflexivity.
+ rewrite <- copy_n_unit by assumption.
+ apply is_unit_sum_Unit in H.
+ rewrite law_neutral_left. reflexivity.
+
+
+ rewrite z0 by auto. rewrite eval_merge_bin. reflexivity.
+ Qed.
+ End sum_correctness.
+ Lemma eval_norm_msets i norm
+ (Comm : Commutative R (Bin.value (e_bin i)))
+ (Hnorm: forall u, eval (norm u) == eval u) : forall h, eval (sum i (norm_msets norm i h)) == eval (sum i h).
Proof.
- intros. unfold sum_to_mset.
- case_eq (is_sum i a).
- intros k Hk. induction n using Pind.
- simpl copy_mset. rewrite (eval_is_sum _ _ Hk). reflexivity.
- rewrite copy_mset_succ. rewrite IHn. rewrite (eval_is_sum _ _ Hk).
- unfold copy at 2. rewrite Prect_succ. reflexivity.
- reflexivity.
- Qed.
+ unfold norm_msets.
+ assert (H :
+ forall h : mset T,
+ eval (sum i (run_msets (norm_msets_ i (is_unit_of i) norm h))) == eval (sum i h) /\ compat_sum_unit (is_unit_of i) (norm_msets_ i (is_unit_of i) norm h)).
+
+ induction h as [[a n] | [a n] h [IHh IHh']]; simpl norm_msets_; split.
+ rewrite z1 by auto. rewrite Hnorm . reflexivity. auto.
+ apply compat_sum_unit_return.
+
+ rewrite z2 by auto. rewrite IHh.
+ rewrite eval_sum_cons. rewrite Hnorm. reflexivity. apply compat_sum_unit_add, IHh'.
+
+ apply H.
+ Defined.
+
+ (** auxiliary lemmas about products *)
+ Section prd_correctness.
+ Variable i : idx.
+ Variable is_unit : idx -> bool.
+ Hypothesis is_unit_prd_Unit : forall j, is_unit j = true-> @Unit X R (Bin.value (e_bin i)) (eval (unit j)).
+
+ Inductive is_prd_spec_ind : T -> @discr (nelist T) -> Prop :=
+ | is_prd_spec_op :
+ forall j l, j = i -> is_prd_spec_ind (prd j l) (Is_op l)
+ | is_prd_spec_unit :
+ forall j, is_unit j = true -> is_prd_spec_ind (unit j) (Is_unit j)
+ | is_prd_spec_nothing :
+ forall u, is_prd_spec_ind u (Is_nothing).
+
+ Lemma is_prd_spec u : is_prd_spec_ind u (is_prd i is_unit u).
+ Proof.
+ unfold is_prd; case u; intros; try constructor.
+ case (eq_idx_spec); intros; subst; try constructor; auto.
+ case_eq (is_unit p); intros; try constructor; auto.
+ Qed.
- (** auxiliary lemmas about products *)
- Lemma prd'_prd i (l: list T): eval (prd' i l) == eval (prd i l).
- Proof. intros i [|? [|? ?]]; reflexivity. Qed.
+ Lemma prd'_prd : forall (l: nelist T), eval (prd' i l) == eval (prd i l).
+ Proof.
+ intros [?|? [|? ?]]; simpl; reflexivity.
+ Qed.
+
+
+ Lemma eval_prd_nil x: eval (prd i (nil x)) == eval x.
+ Proof.
+ rewrite <- prd'_prd. simpl. reflexivity.
+ Qed.
+ Lemma eval_prd_cons a : forall (l: nelist T), eval (prd i (a::l)) == @Bin.value _ _ (e_bin i) (eval a) (eval (prd i l)).
+ Proof.
+ intros [|b l]; simpl; reflexivity.
+ Qed.
+ Lemma eval_prd_app : forall (h k: nelist T), eval (prd i (h++k)) == @Bin.value _ _ (e_bin i) (eval (prd i h)) (eval (prd i k)).
+ Proof.
+ induction h; intro k. simpl; try reflexivity.
+ simpl appne. rewrite 2 eval_prd_cons, IHh, law_assoc. reflexivity.
+ Qed.
+
+ Inductive compat_prd_unit : @m idx (nelist T) -> Prop :=
+ | cpu_left : forall x, is_unit x = true -> compat_prd_unit (left x)
+ | cpu_right : forall m, compat_prd_unit (right m)
+ .
+
+ Lemma compat_prd_unit_return x: compat_prd_unit (return_prd i is_unit x).
+ Proof.
+ unfold return_prd.
+ case (is_prd_spec); intros; try constructor; auto.
+ Qed.
+
+ Lemma compat_prd_unit_add : forall x h,
+ compat_prd_unit h
+ ->
+ compat_prd_unit
+ (add_to_prd i is_unit x
+ h).
+ Proof.
+ intros.
+ unfold add_to_prd.
+ unfold comp.
+ case (is_prd_spec); intros; try constructor; auto.
+ unfold comp; case h; try constructor.
+ unfold comp; case h; try constructor.
+ Qed.
- Lemma eval_prd_cons i a (l: list T): eval (prd i (a::l)) == @Prd.dot _ _ (e_prd i) (eval a) (eval (prd i l)).
- Proof.
- intros i a [|b l]; simpl.
- rewrite dot_neutral_right. reflexivity.
- reflexivity.
- Qed.
+
+ Instance compat_prd_Unit : forall p, compat_prd_unit (left p) ->
+ @Unit X R (Bin.value (e_bin i)) (eval (unit p)).
+ Proof.
+ intros.
+ inversion H; subst. apply is_unit_prd_Unit. assumption.
+ Qed.
+
+ Lemma z0' : forall l (r: @m idx (nelist T)), compat_prd_unit r ->
+ eval (prd i (run_list (comp (@appne T) l r))) == eval (prd i ((appne (l) (run_list r)))).
+ Proof.
+ intros.
+ unfold comp. unfold run_list. case_eq r; intros; auto; subst.
+ rewrite eval_prd_app.
+ rewrite eval_prd_nil.
+ apply compat_prd_Unit in H. rewrite law_neutral_right. reflexivity.
+ reflexivity.
+ Qed.
+
+ Lemma z1' a : eval (prd i (run_list (return_prd i is_unit a))) == eval (prd i (nil a)).
+ Proof.
+ intros. unfold return_prd. unfold run_list.
+ case (is_prd_spec); intros; subst; reflexivity.
+ Qed.
+ Lemma z2' : forall u x, compat_prd_unit x ->
+ eval (prd i ( run_list
+ (add_to_prd i is_unit u x))) == @Bin.value _ _ (e_bin i) (eval u) (eval (prd i (run_list x))).
+ Proof.
+ intros u x Hix.
+ unfold add_to_prd.
+ case (is_prd_spec); intros; subst.
+ rewrite z0' by auto. rewrite eval_prd_app. reflexivity.
+ apply is_unit_prd_Unit in H. rewrite law_neutral_left. reflexivity.
+ rewrite z0' by auto. rewrite eval_prd_app. reflexivity.
+ Qed.
+
+ End prd_correctness.
- Lemma eval_prd_app i (h k: list T): eval (prd i (h++k)) == @Prd.dot _ _ (e_prd i) (eval (prd i h)) (eval (prd i k)).
- Proof.
- induction h; intro k.
- simpl. rewrite dot_neutral_left. reflexivity.
- simpl app. rewrite 2 eval_prd_cons, IHh, dot_assoc. reflexivity.
- Qed.
- Lemma eval_is_prd: forall i (a: T) k, is_prd i a = Some k -> eval (prd i k) = eval a.
- Proof.
- intros i [j l|j l|j l] k; simpl is_prd; try (intro; discriminate).
- case eq_idx_spec; intros ? H.
- injection H. clear H. intros <-. subst. congruence.
- discriminate.
- Qed.
- Lemma eval_prd_to_list: forall i a, eval (prd i (prd_to_list i a)) = eval a.
+
+ Lemma eval_norm_lists i (Hnorm: forall u, eval (norm u) == eval u) : forall h, eval (prd i (norm_lists norm i h)) == eval (prd i h).
Proof.
- intros. unfold prd_to_list.
- case_eq (is_prd i a).
- intros k Hk. apply eval_is_prd. assumption.
- reflexivity.
- Qed.
+ unfold norm_lists.
+ assert (H : forall h : nelist T,
+ eval (prd i (run_list (norm_lists_ i (is_unit_of i) norm h))) ==
+ eval (prd i h)
+ /\ compat_prd_unit (is_unit_of i) (norm_lists_ i (is_unit_of i) norm h)).
+
+
+ induction h as [a | a h [IHh IHh']]; simpl norm_lists_; split.
+ rewrite z1'. simpl. apply Hnorm.
+
+ apply compat_prd_unit_return.
+
+ rewrite z2'. rewrite IHh. rewrite eval_prd_cons. rewrite Hnorm. reflexivity. apply is_unit_of_Unit.
+ auto.
+ apply compat_prd_unit_add. auto.
+ apply H.
+ Defined.
(** correctness of the normalisation function *)
+
Theorem eval_norm: forall u, eval (norm u) == eval u
- with eval_norm_aux: forall i (l: vT i) (f: Sym.type_of i),
+ with eval_norm_aux: forall i (l: vT i) (f: Sym.type_of i),
Proper (@Sym.rel_of X R i) f -> eval_aux (vnorm l) f == eval_aux l f.
Proof.
- induction u.
- simpl norm. rewrite sum'_sum. induction m as [|[u n] m IHm].
- reflexivity.
- simpl merge_map. rewrite eval_merge_sum.
- rewrite eval_sum_to_mset, IHm, eval_norm.
- rewrite eval_sum_cons. reflexivity.
+ induction u as [ p m | p l | ? | ?]; simpl norm.
+ case_eq (is_commutative p); intros.
+ rewrite sum'_sum.
+ apply eval_norm_msets; auto.
+ reflexivity.
- simpl norm. rewrite prd'_prd. induction l.
- reflexivity.
- simpl merge_map.
- rewrite eval_prd_app.
- rewrite eval_prd_to_list, IHl, eval_norm.
- rewrite eval_prd_cons. reflexivity.
+ rewrite prd'_prd.
+ apply eval_norm_lists; auto.
- simpl. apply eval_norm_aux, Sym.morph.
+ apply eval_norm_aux, Sym.morph.
+ reflexivity.
- induction l; simpl; intros f Hf.
- reflexivity.
- rewrite eval_norm. apply IHl, Hf. reflexivity.
+ induction l; simpl; intros f Hf. reflexivity.
+ rewrite eval_norm. apply IHl, Hf; reflexivity.
Qed.
+
(** corollaries, for goal normalisation or decision *)
- Lemma normalise: forall (u v: T), eval (norm u) == eval (norm v) -> eval u == eval v.
+
+ Lemma normalise : forall (u v: T), eval (norm u) == eval (norm v) -> eval u == eval v.
Proof. intros u v. rewrite 2 eval_norm. trivial. Qed.
-
+
Lemma compare_reflect_eq: forall u v, compare u v = Eq -> eval u == eval v.
Proof. intros u v. case (tcompare_weak_spec u v); intros; try congruence. reflexivity. Qed.
Lemma decide: forall (u v: T), compare (norm u) (norm v) = Eq -> eval u == eval v.
Proof. intros u v H. apply normalise. apply compare_reflect_eq. apply H. Qed.
+ Lemma lift_normalise {S} {H : AAC_lift S R}:
+ forall (u v: T), (let x := norm u in let y := norm v in
+ S (eval x) (eval y)) -> S (eval u) (eval v).
+ Proof. destruct H. intros u v; simpl; rewrite 2 eval_norm. trivial. Qed.
+
End s.
+End Internal.
+
+
+(** * Lemmas for performing transitivity steps
+ given an instance of AAC_lift *)
+
+Section t.
+ Context `{AAC_lift}.
+
+ Lemma lift_transitivity_left (y x z : X): E x y -> R y z -> R x z.
+ Proof. destruct H as [Hequiv Hproper]; intros G;rewrite G. trivial. Qed.
+
+ Lemma lift_transitivity_right (y x z : X): E y z -> R x y -> R x z.
+ Proof. destruct H as [Hequiv Hproper]; intros G. rewrite G. trivial. Qed.
+
+ Lemma lift_reflexivity {HR :Reflexive R}: forall x y, E x y -> R x y.
+ Proof. destruct H. intros ? ? G. rewrite G. reflexivity. Qed.
+
+End t.
+
Declare ML Module "aac_tactics".
diff --git a/AAC_coq.ml b/AAC_coq.ml
new file mode 100644
index 0000000..b3ee180
--- /dev/null
+++ b/AAC_coq.ml
@@ -0,0 +1,591 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+(** Interface with Coq *)
+type constr = Term.constr
+
+open Term
+open Names
+open Coqlib
+
+(* The contrib name is used to locate errors when loading constrs *)
+let contrib_name = "aac_tactics"
+
+(* Getting constrs (primitive Coq terms) from existing Coq
+ libraries. *)
+let find_constant contrib dir s =
+ Libnames.constr_of_global (Coqlib.find_reference contrib dir s)
+
+let init_constant dir s = find_constant contrib_name dir s
+
+(* A clause specifying that the [let] should not try to fold anything
+ in the goal *)
+let nowhere =
+ { Tacexpr.onhyps = Some [];
+ Tacexpr.concl_occs = Rawterm.no_occurrences_expr
+ }
+
+let cps_mk_letin
+ (name:string)
+ (c: constr)
+ (k : constr -> Proof_type.tactic)
+: Proof_type.tactic =
+ fun goal ->
+ let name = (Names.id_of_string name) in
+ let name = Tactics.fresh_id [] name goal in
+ let letin = (Tactics.letin_tac None (Name name) c None nowhere) in
+ Tacticals.tclTHEN letin (k (mkVar name)) goal
+
+(** {2 General functions} *)
+
+type goal_sigma = Proof_type.goal Tacmach.sigma
+let goal_update (goal : goal_sigma) evar_map : goal_sigma=
+ let it = Tacmach.sig_it goal in
+ Tacmach.re_sig it evar_map
+
+let fresh_evar goal ty : constr * goal_sigma =
+ let env = Tacmach.pf_env goal in
+ let evar_map = Tacmach.project goal in
+ let (em,x) = Evarutil.new_evar evar_map env ty in
+ x,( goal_update goal em)
+
+let resolve_one_typeclass goal ty : constr*goal_sigma=
+ let env = Tacmach.pf_env goal in
+ let evar_map = Tacmach.project goal in
+ let em,c = Typeclasses.resolve_one_typeclass env evar_map ty in
+ c, (goal_update goal em)
+
+let general_error =
+ "Cannot resolve a typeclass : please report"
+
+let cps_resolve_one_typeclass ?error : Term.types -> (Term.constr -> Proof_type.tactic) -> Proof_type.tactic = fun t k goal ->
+ Tacmach.pf_apply
+ (fun env em -> let em ,c =
+ try Typeclasses.resolve_one_typeclass env em t
+ with Not_found ->
+ begin match error with
+ | None -> Util.anomaly "Cannot resolve a typeclass : please report"
+ | Some x -> Util.error x
+ end
+ in
+ Tacticals.tclTHENLIST [Refiner.tclEVARS em; k c] goal
+ ) goal
+
+
+let nf_evar goal c : Term.constr=
+ let evar_map = Tacmach.project goal in
+ Evarutil.nf_evar evar_map c
+
+let evar_unit (gl : goal_sigma) (x : constr) : constr * goal_sigma =
+ let env = Tacmach.pf_env gl in
+ let evar_map = Tacmach.project gl in
+ let (em,x) = Evarutil.new_evar evar_map env x in
+ x,(goal_update gl em)
+
+let evar_binary (gl: goal_sigma) (x : constr) =
+ let env = Tacmach.pf_env gl in
+ let evar_map = Tacmach.project gl in
+ let ty = mkArrow x (mkArrow x x) in
+ let (em,x) = Evarutil.new_evar evar_map env ty in
+ x,( goal_update gl em)
+
+let evar_relation (gl: goal_sigma) (x: constr) =
+ let env = Tacmach.pf_env gl in
+ let evar_map = Tacmach.project gl in
+ let ty = mkArrow x (mkArrow x (mkSort prop_sort)) in
+ let (em,r) = Evarutil.new_evar evar_map env ty in
+ r,( goal_update gl em)
+
+let cps_evar_relation (x: constr) k = fun goal ->
+ Tacmach.pf_apply
+ (fun env em ->
+ let ty = mkArrow x (mkArrow x (mkSort prop_sort)) in
+ let (em,r) = Evarutil.new_evar em env ty in
+ Tacticals.tclTHENLIST [Refiner.tclEVARS em; k r] goal
+ ) goal
+
+(* decomp_term : constr -> (constr, types) kind_of_term *)
+let decomp_term c = kind_of_term (strip_outer_cast c)
+
+let lapp c v = mkApp (Lazy.force c, v)
+
+(** {2 Bindings with Coq' Standard Library} *)
+module Std = struct
+(* Here we package the module to be able to use List, later *)
+
+module Pair = struct
+
+ let path = ["Coq"; "Init"; "Datatypes"]
+ let typ = lazy (init_constant path "prod")
+ let pair = lazy (init_constant path "pair")
+ let of_pair t1 t2 (x,y) =
+ mkApp (Lazy.force pair, [| t1; t2; x ; y|] )
+end
+
+module Bool = struct
+ let path = ["Coq"; "Init"; "Datatypes"]
+ let typ = lazy (init_constant path "bool")
+ let ctrue = lazy (init_constant path "true")
+ let cfalse = lazy (init_constant path "false")
+ let of_bool b =
+ if b then Lazy.force ctrue else Lazy.force cfalse
+end
+
+module Comparison = struct
+ let path = ["Coq"; "Init"; "Datatypes"]
+ let typ = lazy (init_constant path "comparison")
+ let eq = lazy (init_constant path "Eq")
+ let lt = lazy (init_constant path "Lt")
+ let gt = lazy (init_constant path "Gt")
+end
+
+module Leibniz = struct
+ let path = ["Coq"; "Init"; "Logic"]
+ let eq_refl = lazy (init_constant path "eq_refl")
+ let eq_refl ty x = lapp eq_refl [| ty;x|]
+end
+
+module Option = struct
+ let path = ["Coq"; "Init"; "Datatypes"]
+ let typ = lazy (init_constant path "option")
+ let some = lazy (init_constant path "Some")
+ let none = lazy (init_constant path "None")
+ let some t x = mkApp (Lazy.force some, [| t ; x|])
+ let none t = mkApp (Lazy.force none, [| t |])
+ let of_option t x = match x with
+ | Some x -> some t x
+ | None -> none t
+end
+
+module Pos = struct
+
+ let path = ["Coq" ; "NArith"; "BinPos"]
+ let typ = lazy (init_constant path "positive")
+ let xI = lazy (init_constant path "xI")
+ let xO = lazy (init_constant path "xO")
+ let xH = lazy (init_constant path "xH")
+
+ (* A coq positive from an int *)
+ let of_int n =
+ let rec aux n =
+ begin match n with
+ | n when n < 1 -> assert false
+ | 1 -> Lazy.force xH
+ | n -> mkApp
+ (
+ (if n mod 2 = 0
+ then Lazy.force xO
+ else Lazy.force xI
+ ), [| aux (n/2)|]
+ )
+ end
+ in
+ aux n
+end
+
+module Nat = struct
+ let path = ["Coq" ; "Init"; "Datatypes"]
+ let typ = lazy (init_constant path "nat")
+ let _S = lazy (init_constant path "S")
+ let _O = lazy (init_constant path "O")
+ (* A coq nat from an int *)
+ let of_int n =
+ let rec aux n =
+ begin match n with
+ | n when n < 0 -> assert false
+ | 0 -> Lazy.force _O
+ | n -> mkApp
+ (
+ (Lazy.force _S
+ ), [| aux (n-1)|]
+ )
+ end
+ in
+ aux n
+end
+
+(** Lists from the standard library*)
+module List = struct
+ let path = ["Coq"; "Lists"; "List"]
+ let typ = lazy (init_constant path "list")
+ let nil = lazy (init_constant path "nil")
+ let cons = lazy (init_constant path "cons")
+ let cons ty h t =
+ mkApp (Lazy.force cons, [| ty; h ; t |])
+ let nil ty =
+ (mkApp (Lazy.force nil, [| ty |]))
+ let rec of_list ty = function
+ | [] -> nil ty
+ | t::q -> cons ty t (of_list ty q)
+ let type_of_list ty =
+ mkApp (Lazy.force typ, [|ty|])
+end
+
+(** Morphisms *)
+module Classes =
+struct
+ let classes_path = ["Coq";"Classes"; ]
+ let morphism = lazy (init_constant (classes_path@["Morphisms"]) "Proper")
+ let equivalence = lazy (init_constant (classes_path@ ["RelationClasses"]) "Equivalence")
+ let reflexive = lazy (init_constant (classes_path@ ["RelationClasses"]) "Reflexive")
+ let transitive = lazy (init_constant (classes_path@ ["RelationClasses"]) "Transitive")
+
+ (** build the type [Equivalence ty rel] *)
+ let mk_equivalence ty rel =
+ mkApp (Lazy.force equivalence, [| ty; rel|])
+
+
+ (** build the type [Reflexive ty rel] *)
+ let mk_reflexive ty rel =
+ mkApp (Lazy.force reflexive, [| ty; rel|])
+
+ (** build the type [Proper rel t] *)
+ let mk_morphism ty rel t =
+ mkApp (Lazy.force morphism, [| ty; rel; t|])
+
+ (** build the type [Transitive ty rel] *)
+ let mk_transitive ty rel =
+ mkApp (Lazy.force transitive, [| ty; rel|])
+end
+
+module Relation = struct
+ type t =
+ {
+ carrier : constr;
+ r : constr;
+ }
+
+ let make ty r = {carrier = ty; r = r }
+ let split t = t.carrier, t.r
+end
+
+module Transitive = struct
+ type t =
+ {
+ carrier : constr;
+ leq : constr;
+ transitive : constr;
+ }
+ let make ty leq transitive = {carrier = ty; leq = leq; transitive = transitive}
+ let infer goal ty leq =
+ let ask = Classes.mk_transitive ty leq in
+ let transitive , goal = resolve_one_typeclass goal ask in
+ make ty leq transitive, goal
+ let from_relation goal rlt =
+ infer goal (rlt.Relation.carrier) (rlt.Relation.r)
+ let cps_from_relation rlt k =
+ let ty =rlt.Relation.carrier in
+ let r = rlt.Relation.r in
+ let ask = Classes.mk_transitive ty r in
+ cps_resolve_one_typeclass ask
+ (fun trans -> k (make ty r trans) )
+ let to_relation t =
+ {Relation.carrier = t.carrier; Relation.r = t.leq}
+
+end
+
+module Equivalence = struct
+ type t =
+ {
+ carrier : constr;
+ eq : constr;
+ equivalence : constr;
+ }
+ let make ty eq equivalence = {carrier = ty; eq = eq; equivalence = equivalence}
+ let infer goal ty eq =
+ let ask = Classes.mk_equivalence ty eq in
+ let equivalence , goal = resolve_one_typeclass goal ask in
+ make ty eq equivalence, goal
+ let from_relation goal rlt =
+ infer goal (rlt.Relation.carrier) (rlt.Relation.r)
+ let cps_from_relation rlt k =
+ let ty =rlt.Relation.carrier in
+ let r = rlt.Relation.r in
+ let ask = Classes.mk_equivalence ty r in
+ cps_resolve_one_typeclass ask (fun equiv -> k (make ty r equiv) )
+ let to_relation t =
+ {Relation.carrier = t.carrier; Relation.r = t.eq}
+ let split t =
+ t.carrier, t.eq, t.equivalence
+end
+end
+(**[ match_as_equation goal eqt] see [eqt] as an equation. An
+ optionnal rel_context can be provided to ensure taht the term
+ remains typable*)
+let match_as_equation ?(context = Term.empty_rel_context) goal equation : (constr*constr* Std.Relation.t) option =
+ let env = Tacmach.pf_env goal in
+ let env = Environ.push_rel_context context env in
+ let evar_map = Tacmach.project goal in
+ begin
+ match decomp_term equation with
+ | App(c,ca) when Array.length ca >= 2 ->
+ let n = Array.length ca in
+ let left = ca.(n-2) in
+ let right = ca.(n-1) in
+ let r = (mkApp (c, Array.sub ca 0 (n - 2))) in
+ let carrier = Typing.type_of env evar_map left in
+ let rlt =Std.Relation.make carrier r
+ in
+ Some (left, right, rlt )
+ | _ -> None
+ end
+
+
+(** {1 Tacticals} *)
+
+let tclTIME msg tac = fun gl ->
+ let u = Sys.time () in
+ let r = tac gl in
+ let _ = Pp.msgnl (Pp.str (Printf.sprintf "%s:%fs" msg (Sys.time ()-. u))) in
+ r
+
+let tclDEBUG msg tac = fun gl ->
+ let _ = Pp.msgnl (Pp.str msg) in
+ tac gl
+
+let tclPRINT tac = fun gl ->
+ let _ = Pp.msgnl (Printer.pr_constr (Tacmach.pf_concl gl)) in
+ tac gl
+
+
+(** {1 Error related mechanisms} *)
+(* functions to handle the failures of our tactic. Some should be
+ reported [anomaly], some are on behalf of the user [user_error]*)
+let anomaly msg =
+ Util.anomaly ("[aac_tactics] " ^ msg)
+
+let user_error msg =
+ Util.error ("[aac_tactics] " ^ msg)
+
+let warning msg =
+ Pp.msg_warning (Pp.str ("[aac_tactics]" ^ msg))
+
+
+(** {1 Rewriting tactics used in aac_rewrite} *)
+module Rewrite = struct
+(** Some informations about the hypothesis, with an (informal)
+ invariant:
+ - [typeof hyp = hyptype]
+ - [hyptype = forall context, body]
+ - [body = rel left right]
+
+*)
+
+type hypinfo =
+ {
+ hyp : Term.constr; (** the actual constr corresponding to the hypothese *)
+ hyptype : Term.constr; (** the type of the hypothesis *)
+ context : Term.rel_context; (** the quantifications of the hypothese *)
+ body : Term.constr; (** the body of the type of the hypothesis*)
+ rel : Std.Relation.t; (** the relation *)
+ left : Term.constr; (** left hand side *)
+ right : Term.constr; (** right hand side *)
+ l2r : bool; (** rewriting from left to right *)
+ }
+
+let get_hypinfo c ~l2r ?check_type (k : hypinfo -> Proof_type.tactic) : Proof_type.tactic = fun goal ->
+ let ctype = Tacmach.pf_type_of goal c in
+ let (rel_context, body_type) = Term.decompose_prod_assum ctype in
+ let rec check f e =
+ match decomp_term e with
+ | Term.Rel i ->
+ let name, constr_option, types = Term.lookup_rel i rel_context
+ in f types
+ | _ -> Term.fold_constr (fun acc x -> acc && check f x) true e
+ in
+ begin match check_type with
+ | None -> ()
+ | Some f ->
+ if not (check f body_type)
+ then user_error "Unable to deal with higher-order or heterogeneous patterns";
+ end;
+ begin
+ match match_as_equation ~context:rel_context goal body_type with
+ | None ->
+ user_error "The hypothesis is not an applied relation"
+ | Some (hleft,hright,hrlt) ->
+ k {
+ hyp = c;
+ hyptype = ctype;
+ body = body_type;
+ l2r = l2r;
+ context = rel_context;
+ rel = hrlt ;
+ left =hleft;
+ right = hright;
+ }
+ goal
+ end
+
+
+(* The problem : Given a term to rewrite of type [H :forall xn ... x1,
+ t], we have to instanciate the subset of [xi] of type
+ [carrier]. [subst : (int * constr)] is the mapping the debruijn
+ indices in [t] to the [constrs]. We need to handle the rest of the
+ indexes. Two ways :
+
+ - either using fresh evars and rewriting [H tn ?1 tn-2 ?2 ]
+ - either building a term like [fun 1 2 => H tn 1 tn-2 2]
+
+ Both these terms have the same type.
+*)
+
+
+(* Fresh evars for everyone (should be the good way to do this
+ recompose in Coq v8.4) *)
+let recompose_prod
+ (context : Term.rel_context)
+ (subst : (int * Term.constr) list)
+ env
+ em
+ : Evd.evar_map * Term.constr list =
+ (* the last index of rel relevant for the rewriting *)
+ let min_n = List.fold_left
+ (fun acc (x,_) -> min acc x)
+ (List.length context) subst in
+ let rec aux context acc em n =
+ let _ = Printf.printf "%i\n%!" n in
+ match context with
+ | [] ->
+ env, em, acc
+ | ((name,c,ty) as t)::q ->
+ let env, em, acc = aux q acc em (n+1) in
+ if n >= min_n
+ then
+ let em,x =
+ try em, List.assoc n subst
+ with | Not_found ->
+ Evarutil.new_evar em env (Term.substl acc ty)
+ in
+ (Environ.push_rel t env), em,x::acc
+ else
+ env,em,acc
+ in
+ let _,em,acc = aux context [] em 1 in
+ em, acc
+
+(* no fresh evars : instead, use a lambda abstraction around an
+ application to handle non-instanciated variables. *)
+
+let recompose_prod'
+ (context : Term.rel_context)
+ (subst : (int *Term.constr) list)
+ c
+ =
+ let rec popn pop n l =
+ if n <= 0 then l
+ else match l with
+ | [] -> []
+ | t::q -> pop t :: (popn pop (n-1) q)
+ in
+ let pop_rel_decl (name,c,ty) = (name,c,Termops.pop ty) in
+ let rec aux sign n next app ctxt =
+ match sign with
+ | [] -> List.rev app, List.rev ctxt
+ | decl::sign ->
+ try let term = (List.assoc n subst) in
+ aux sign (n+1) next (term::app) (None :: ctxt)
+ with
+ | Not_found ->
+ let term = Term.mkRel next in
+ aux sign (n+1) (next+1) (term::app) (Some decl :: ctxt)
+ in
+ let app,ctxt = aux context 1 1 [] [] in
+ (* substitutes in the context *)
+ let rec update ctxt app =
+ match ctxt,app with
+ | [],_ -> []
+ | _,[] -> []
+ | None :: sign, _ :: app ->
+ None :: update sign (List.map (Termops.pop) app)
+ | Some decl :: sign, _ :: app ->
+ Some (Term.substl_decl app decl)::update sign (List.map (Termops.pop) app)
+ in
+ let ctxt = update ctxt app in
+ (* updates the rel accordingly, taking some care not to go to far
+ beyond: it is important not to lift indexes homogeneously, we
+ have to update *)
+ let rec update ctxt res n =
+ match ctxt with
+ | [] -> List.rev res
+ | None :: sign ->
+ (update (sign) (popn pop_rel_decl n res) 0)
+ | Some decl :: sign ->
+ update sign (decl :: res) (n+1)
+ in
+ let ctxt = update ctxt [] 0 in
+ let c = Term.applistc c (List.rev app) in
+ let c = Term.it_mkLambda_or_LetIn c (ctxt) in
+ c
+
+(* Version de Matthieu
+let subst_rel_context k cstr ctx =
+ let (_, ctx') =
+ List.fold_left (fun (k, ctx') (id, b, t) -> (succ k, (id, Option.map
+ (Term.substnl [cstr] k) b, Term.substnl [cstr] k t) :: ctx')) (k, [])
+ ctx in List.rev ctx'
+
+let recompose_prod' context subst c =
+ let len = List.length context in
+ let rec aux sign n next app ctxt =
+ match sign with
+ | [] -> List.rev app, List.rev ctxt
+ | decl::sign ->
+ try let term = (List.assoc n subst) in
+ aux (subst_rel_context 0 term sign) (pred n) (succ next)
+ (term::List.map (Term.lift (-1)) app) ctxt
+ with Not_found ->
+ let term = Term.mkRel (len - next) in
+ aux sign (pred n) (succ next) (term::app) (decl :: ctxt)
+ in
+ let app,ctxt = aux (List.rev context) len 0 [] [] in
+ Term.it_mkLambda_or_LetIn (Term.applistc c(app)) (List.rev ctxt)
+*)
+
+let build
+ (h : hypinfo)
+ (subst : (int *Term.constr) list)
+ (k :Term.constr -> Proof_type.tactic)
+ : Proof_type.tactic = fun goal ->
+ let c = recompose_prod' h.context subst h.hyp in
+ Tacticals.tclTHENLIST [k c] goal
+
+let build_with_evar
+ (h : hypinfo)
+ (subst : (int *Term.constr) list)
+ (k :Term.constr -> Proof_type.tactic)
+ : Proof_type.tactic
+ = fun goal ->
+ Tacmach.pf_apply
+ (fun env em ->
+ let evar_map, acc = recompose_prod h.context subst env em in
+ let c = Term.applistc h.hyp (List.rev acc) in
+ Tacticals.tclTHENLIST [Refiner.tclEVARS evar_map; k c] goal
+ ) goal
+
+
+let rewrite ?(abort=false)hypinfo subst k =
+ build hypinfo subst
+ (fun rew ->
+ let tac =
+ if not abort then
+ Equality.general_rewrite_bindings
+ hypinfo.l2r
+ Termops.all_occurrences
+ false
+ (rew,Rawterm.NoBindings)
+ true
+ else
+ Tacticals.tclIDTAC
+ in k tac
+ )
+
+
+end
+
+include Std
diff --git a/AAC_coq.mli b/AAC_coq.mli
new file mode 100644
index 0000000..a0f2ce0
--- /dev/null
+++ b/AAC_coq.mli
@@ -0,0 +1,231 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+(** Interface with Coq where we define some handlers for Coq's API,
+ and we import several definitions from Coq's standard library.
+
+ This general purpose library could be reused by other plugins.
+
+ Some salient points:
+ - we use Caml's module system to mimic Coq's one, and avoid cluttering
+ the namespace;
+ - we also provide several handlers for standard coq tactics, in
+ order to provide a unified setting (we replace functions that
+ modify the evar_map by functions that operate on the whole
+ goal, and repack the modified evar_map with it).
+
+*)
+
+(** {2 Getting Coq terms from the environment} *)
+
+val init_constant : string list -> string -> Term.constr
+
+(** {2 General purpose functions} *)
+
+type goal_sigma = Proof_type.goal Tacmach.sigma
+val goal_update : goal_sigma -> Evd.evar_map -> goal_sigma
+val resolve_one_typeclass : Proof_type.goal Tacmach.sigma -> Term.types -> Term.constr * goal_sigma
+val cps_resolve_one_typeclass: ?error:string -> Term.types -> (Term.constr -> Proof_type.tactic) -> Proof_type.tactic
+val nf_evar : goal_sigma -> Term.constr -> Term.constr
+val fresh_evar :goal_sigma -> Term.types -> Term.constr* goal_sigma
+val evar_unit :goal_sigma ->Term.constr -> Term.constr* goal_sigma
+val evar_binary: goal_sigma -> Term.constr -> Term.constr* goal_sigma
+val evar_relation: goal_sigma -> Term.constr -> Term.constr* goal_sigma
+val cps_evar_relation : Term.constr -> (Term.constr -> Proof_type.tactic) -> Proof_type.tactic
+(** [cps_mk_letin name v] binds the constr [v] using a letin tactic *)
+val cps_mk_letin : string -> Term.constr -> ( Term.constr -> Proof_type.tactic) -> Proof_type.tactic
+
+
+val decomp_term : Term.constr -> (Term.constr , Term.types) Term.kind_of_term
+val lapp : Term.constr lazy_t -> Term.constr array -> Term.constr
+
+(** {2 Bindings with Coq' Standard Library} *)
+
+(** Coq lists *)
+module List:
+sig
+ (** [of_list ty l] *)
+ val of_list:Term.constr ->Term.constr list ->Term.constr
+
+ (** [type_of_list ty] *)
+ val type_of_list:Term.constr ->Term.constr
+end
+
+(** Coq pairs *)
+module Pair:
+sig
+ val typ:Term.constr lazy_t
+ val pair:Term.constr lazy_t
+ val of_pair : Term.constr -> Term.constr -> Term.constr * Term.constr -> Term.constr
+end
+
+module Bool : sig
+ val typ : Term.constr lazy_t
+ val of_bool : bool -> Term.constr
+end
+
+
+module Comparison : sig
+ val typ : Term.constr lazy_t
+ val eq : Term.constr lazy_t
+ val lt : Term.constr lazy_t
+ val gt : Term.constr lazy_t
+end
+
+module Leibniz : sig
+ val eq_refl : Term.types -> Term.constr -> Term.constr
+end
+
+module Option : sig
+ val some : Term.constr -> Term.constr -> Term.constr
+ val none : Term.constr -> Term.constr
+ val of_option : Term.constr -> Term.constr option -> Term.constr
+end
+
+(** Coq positive numbers (pos) *)
+module Pos:
+sig
+ val typ:Term.constr lazy_t
+ val of_int: int ->Term.constr
+end
+
+(** Coq unary numbers (peano) *)
+module Nat:
+sig
+ val typ:Term.constr lazy_t
+ val of_int: int ->Term.constr
+end
+
+(** Coq typeclasses *)
+module Classes:
+sig
+ val mk_morphism: Term.constr -> Term.constr -> Term.constr -> Term.constr
+ val mk_equivalence: Term.constr -> Term.constr -> Term.constr
+ val mk_reflexive: Term.constr -> Term.constr -> Term.constr
+ val mk_transitive: Term.constr -> Term.constr -> Term.constr
+end
+
+module Relation : sig
+ type t = {carrier : Term.constr; r : Term.constr;}
+ val make : Term.constr -> Term.constr -> t
+ val split : t -> Term.constr * Term.constr
+end
+
+module Transitive : sig
+ type t =
+ {
+ carrier : Term.constr;
+ leq : Term.constr;
+ transitive : Term.constr;
+ }
+ val make : Term.constr -> Term.constr -> Term.constr -> t
+ val infer : goal_sigma -> Term.constr -> Term.constr -> t * goal_sigma
+ val from_relation : goal_sigma -> Relation.t -> t * goal_sigma
+ val cps_from_relation : Relation.t -> (t -> Proof_type.tactic) -> Proof_type.tactic
+ val to_relation : t -> Relation.t
+end
+
+module Equivalence : sig
+ type t =
+ {
+ carrier : Term.constr;
+ eq : Term.constr;
+ equivalence : Term.constr;
+ }
+ val make : Term.constr -> Term.constr -> Term.constr -> t
+ val infer : goal_sigma -> Term.constr -> Term.constr -> t * goal_sigma
+ val from_relation : goal_sigma -> Relation.t -> t * goal_sigma
+ val cps_from_relation : Relation.t -> (t -> Proof_type.tactic) -> Proof_type.tactic
+ val to_relation : t -> Relation.t
+ val split : t -> Term.constr * Term.constr * Term.constr
+end
+
+(** [match_as_equation ?context goal c] try to decompose c as a
+ relation applied to two terms. An optionnal rel_context can be
+ provided to ensure that the term remains typable *)
+val match_as_equation : ?context:Term.rel_context -> goal_sigma -> Term.constr -> (Term.constr * Term.constr * Relation.t) option
+
+(** {2 Some tacticials} *)
+
+(** time the execution of a tactic *)
+val tclTIME : string -> Proof_type.tactic -> Proof_type.tactic
+
+(** emit debug messages to see which tactics are failing *)
+val tclDEBUG : string -> Proof_type.tactic -> Proof_type.tactic
+
+(** print the current goal *)
+val tclPRINT : Proof_type.tactic -> Proof_type.tactic
+
+
+(** {2 Error related mechanisms} *)
+
+val anomaly : string -> 'a
+val user_error : string -> 'a
+val warning : string -> unit
+
+
+(** {2 Rewriting tactics used in aac_rewrite} *)
+
+module Rewrite : sig
+
+(** The rewriting tactics used in aac_rewrite, build as handlers of the usual [setoid_rewrite]*)
+
+
+(** {2 Datatypes} *)
+
+(** We keep some informations about the hypothesis, with an (informal)
+ invariant:
+ - [typeof hyp = typ]
+ - [typ = forall context, body]
+ - [body = rel left right]
+
+*)
+type hypinfo =
+ {
+ hyp : Term.constr; (** the actual constr corresponding to the hypothese *)
+ hyptype : Term.constr; (** the type of the hypothesis *)
+ context : Term.rel_context; (** the quantifications of the hypothese *)
+ body : Term.constr; (** the body of the hypothese*)
+ rel : Relation.t; (** the relation *)
+ left : Term.constr; (** left hand side *)
+ right : Term.constr; (** right hand side *)
+ l2r : bool; (** rewriting from left to right *)
+ }
+
+(** [get_hypinfo H l2r ?check_type k] analyse the hypothesis H, and
+ build the related hypinfo, in CPS style. Moreover, an optionnal
+ function can be provided to check the type of every free
+ variable of the body of the hypothesis. *)
+val get_hypinfo :Term.constr -> l2r:bool -> ?check_type:(Term.types -> bool) -> (hypinfo -> Proof_type.tactic) -> Proof_type.tactic
+
+(** {2 Rewriting with bindings}
+
+ The problem : Given a term to rewrite of type [H :forall xn ... x1,
+ t], we have to instanciate the subset of [xi] of type
+ [carrier]. [subst : (int * constr)] is the mapping the De Bruijn
+ indices in [t] to the [constrs]. We need to handle the rest of the
+ indexes. Two ways :
+
+ - either using fresh evars and rewriting [H tn ?1 tn-2 ?2 ]
+ - either building a term like [fun 1 2 => H tn 1 tn-2 2]
+
+ Both these terms have the same type.
+*)
+
+(** build the constr to rewrite, in CPS style, with lambda abstractions *)
+val build : hypinfo -> (int * Term.constr) list -> (Term.constr -> Proof_type.tactic) -> Proof_type.tactic
+
+(** build the constr to rewrite, in CPS style, with evars *)
+val build_with_evar : hypinfo -> (int * Term.constr) list -> (Term.constr -> Proof_type.tactic) -> Proof_type.tactic
+
+(** [rewrite ?abort hypinfo subst] builds the rewriting tactic
+ associated with the given [subst] and [hypinfo], and feeds it to
+ the given continuation. If [abort] is set to [true], we build
+ [tclIDTAC] instead. *)
+val rewrite : ?abort:bool -> hypinfo -> (int * Term.constr) list -> (Proof_type.tactic -> Proof_type.tactic) -> Proof_type.tactic
+end
diff --git a/AAC_helper.ml b/AAC_helper.ml
new file mode 100644
index 0000000..637def1
--- /dev/null
+++ b/AAC_helper.ml
@@ -0,0 +1,41 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+module type CONTROL = sig
+ val debug : bool
+ val time : bool
+ val printing : bool
+end
+
+module Debug (X : CONTROL) = struct
+ open X
+ let debug x =
+ if debug then
+ Printf.printf "%s\n%!" x
+
+
+ let time f x fmt =
+ if time then
+ let t = Sys.time () in
+ let r = f x in
+ Printf.printf fmt (Sys.time () -. t);
+ r
+ else f x
+
+ let pr_constr msg constr =
+ if printing then
+ ( Pp.msgnl (Pp.str (Printf.sprintf "=====%s====" msg));
+ Pp.msgnl (Printer.pr_constr constr);
+ )
+
+
+ let debug_exception msg e =
+ debug (msg ^ (Printexc.to_string e))
+
+
+end
diff --git a/AAC_helper.mli b/AAC_helper.mli
new file mode 100644
index 0000000..f4e4454
--- /dev/null
+++ b/AAC_helper.mli
@@ -0,0 +1,33 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+(** Debugging functions, that can be triggered on a per-file base. *)
+
+module type CONTROL =
+sig
+ val debug : bool
+ val time : bool
+ val printing : bool
+end
+module Debug :
+ functor (X : CONTROL) ->
+sig
+ (** {!debug} prints the string and end it with a newline *)
+ val debug : string -> unit
+ val debug_exception : string -> exn -> unit
+
+ (** {!time} computes the time spent in a function, and then
+ print it using the given format *)
+ val time :
+ ('a -> 'b) -> 'a -> (float -> unit, out_channel, unit) format -> 'b
+
+ (** {!pr_constr} print a Coq constructor, that can be labelled
+ by a string *)
+ val pr_constr : string -> Term.constr -> unit
+
+ end
diff --git a/AAC_matcher.ml b/AAC_matcher.ml
new file mode 100644
index 0000000..47ab840
--- /dev/null
+++ b/AAC_matcher.ml
@@ -0,0 +1,1206 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+(** This module defines our matching functions, modulo associativity
+ and commutativity (AAC).
+
+ The basic idea is to find a substitution [env] such that the
+ pattern [p] instantiated by [env] is equal to [t] modulo AAC.
+
+ We proceed by structural decomposition of the pattern, and try all
+ possible non-deterministic split of the subject when needed. The
+ function {!matcher} is limited to top-level matching, that is, the
+ subject must make a perfect match against the pattern ([x+x] do
+ not match [a+a+b] ). We use a search monad {!Search} to perform
+ non-deterministic splits in an almost transparent way. We also
+ provide a function {!subterm} for finding a match that is a
+ subterm modulo AAC of the subject. Therefore, we are able to solve
+ the aforementioned case [x+x] against [a+b+a].
+
+ This file is structured as follows. First, we define the search
+ monad. Then,we define the two representations of terms (one
+ representing the AST, and one in normal form ), and environments
+ from variables to terms. Then, we use these parts to solve
+ matching problem. Finally, we wrap this function {!matcher} into
+ {!subterm}
+*)
+
+
+module Control =
+struct
+ let debug = false
+ let time = false
+ let printing = false
+end
+
+module Debug = AAC_helper.Debug (Control)
+open Debug
+
+type symbol = int
+type var = int
+type units = (symbol * symbol) list (* from AC/A symbols to the unit *)
+type ext_units =
+ {
+ unit_for : units;
+ is_ac : (symbol * bool) list
+ }
+
+let print_units units=
+ List.iter (fun (op,unit) -> Printf.printf "%i %i\t" op unit) units;
+ Printf.printf "\n%!"
+
+exception NoUnit
+
+let get_unit units op =
+ try List.assoc op units
+ with
+ | Not_found -> raise NoUnit
+
+let is_unit units op unit = List.mem (op,unit) units
+
+
+open AAC_search_monad
+
+type 'a mset = ('a * int) list
+let linear p =
+ let rec ncons t l = function
+ | 0 -> l
+ | n -> t::ncons t l (n-1)
+ in
+ let rec aux = function
+ [ ] -> []
+ | (t,n)::q -> let q = aux q in
+ ncons t q n
+ in aux p
+
+
+
+(** The module {!Terms} defines two different types for expressions.
+
+ - a public type {!Terms.t} that represent abstract syntax trees of
+ expressions with binary associative (and commutative) operators
+
+ - a private type {!Terms.nf_term} that represent an equivalence
+ class for terms that are equal modulo AAC. The constructions
+ functions on this type ensure the property that the term is in
+ normal form (that is, no sum can appear as a subterm of the same
+ sum, no trailing units, etc...).
+
+*)
+
+module Terms : sig
+
+ (** {1 Abstract syntax tree of terms}
+
+ Terms represented using this datatype are representation of the
+ AST of an expression. *)
+
+ type t =
+ Dot of (symbol * t * t)
+ | Plus of (symbol * t * t)
+ | Sym of (symbol * t array)
+ | Var of var
+ | Unit of symbol
+
+ val equal_aac : units -> t -> t -> bool
+ val size: t -> int
+
+ (** {1 Terms in normal form}
+
+ A term in normal form is the canonical representative of the
+ equivalence class of all the terms that are equal modulo
+ Associativity and Commutativity. Outside the {!AAC_matcher} module,
+ one does not need to access the actual representation of this
+ type. *)
+
+
+ type nf_term = private
+ | TAC of symbol * nf_term mset
+ | TA of symbol * nf_term list
+ | TSym of symbol * nf_term list
+ | TUnit of symbol
+ | TVar of var
+
+
+ (** {2 Constructors: we ensure that the terms are always
+ normalised
+
+ braibant - Fri 27 Aug 2010, 15:11
+ Moreover, we assure that we will not build empty sums or empty
+ products, leaving this task to the caller function.
+ }
+ *)
+ val mk_TAC : units -> symbol -> nf_term mset -> nf_term
+ val mk_TA : units -> symbol -> nf_term list -> nf_term
+ val mk_TSym : symbol -> nf_term list -> nf_term
+ val mk_TVar : var -> nf_term
+ val mk_TUnit : symbol -> nf_term
+
+ (** {2 Comparisons} *)
+ val nf_term_compare : nf_term -> nf_term -> int
+ val nf_equal : nf_term -> nf_term -> bool
+
+ (** {2 Printing function} *)
+ val sprint_nf_term : nf_term -> string
+
+ (** {2 Conversion functions} *)
+ val term_of_t : units -> t -> nf_term
+ val t_of_term : nf_term -> t (* we could return the units here *)
+end
+ = struct
+
+ type t =
+ Dot of (symbol * t * t)
+ | Plus of (symbol * t * t)
+ | Sym of (symbol * t array)
+ | Var of var
+ | Unit of symbol
+
+ let rec size = function
+ | Dot (_,x,y)
+ | Plus (_,x,y) -> size x+ size y + 1
+ | Sym (_,v)-> Array.fold_left (fun acc x -> size x + acc) 1 v
+ | _ -> 1
+
+
+ type nf_term =
+ | TAC of symbol * nf_term mset
+ | TA of symbol * nf_term list
+ | TSym of symbol * nf_term list
+ | TUnit of symbol
+ | TVar of var
+
+
+
+
+ (** {2 Constructors: we ensure that the terms are always
+ normalised} *)
+
+ (** {3 Pre constructors : These constructors ensure that sums and
+ products are not degenerated (no trailing units)} *)
+ let mk_TAC' units (s : symbol) l = match l with
+ | [] -> TUnit (get_unit units s )
+ | [_,0] -> assert false
+ | [t,1] -> t
+ | _ -> TAC (s,l)
+ let mk_TA' units (s : symbol) l = match l with
+ | [] -> TUnit (get_unit units s )
+ | [t] -> t
+ | _ -> TA (s,l)
+
+
+ (** {2 Comparison} *)
+
+ let nf_term_compare = Pervasives.compare
+ let nf_equal a b =
+ a = b
+
+ (** [merge_ac comp l1 l2] merges two lists of terms with coefficients
+ into one. Terms that are equal modulo the comparison function
+ [comp] will see their arities added. *)
+
+ (* This function could be improved by the use of sorted msets *)
+ let merge_ac (compare : 'a -> 'a -> int) (l : 'a mset) (l' : 'a mset) : 'a mset =
+ let rec aux l l'=
+ match l,l' with
+ | [], _ -> l'
+ | _, [] -> l
+ | (t,tar)::q, (t',tar')::q' ->
+ begin match compare t t' with
+ | 0 -> ( t,tar+tar'):: aux q q'
+ | -1 -> (t, tar):: aux q l'
+ | _ -> (t', tar'):: aux l q'
+ end
+ in aux l l'
+
+ (** [merge_map f l] uses the combinator [f] to combine the head of the
+ list [l] with the merge_maped tail of [l] *)
+ let rec merge_map (f : 'a -> 'b list -> 'b list) (l : 'a list) : 'b list =
+ match l with
+ | [] -> []
+ | t::q -> f t (merge_map f q)
+
+
+ (** This function has to deal with the arities *)
+ let rec merge (l : nf_term mset) (l' : nf_term mset) : nf_term mset=
+ merge_ac nf_term_compare l l'
+
+ let extract_A units s t =
+ match t with
+ | TA (s',l) when s' = s -> l
+ | TUnit u when is_unit units s u -> []
+ | _ -> [t]
+
+ let extract_AC units s (t,ar) : nf_term mset =
+ match t with
+ | TAC (s',l) when s' = s -> List.map (fun (x,y) -> (x,y*ar)) l
+ | TUnit u when is_unit units s u -> []
+ | _ -> [t,ar]
+
+
+ (** {3 Constructors of {!nf_term}}*)
+ let mk_TAC units (s : symbol) (l : (nf_term *int) list) =
+ mk_TAC' units s
+ (merge_map (fun u l -> merge (extract_AC units s u) l) l)
+ let mk_TA units s l =
+ mk_TA' units s
+ (merge_map (fun u l -> (extract_A units s u) @ l) l)
+ let mk_TSym s l = TSym (s,l)
+ let mk_TVar v = TVar v
+ let mk_TUnit s = TUnit s
+
+
+ (** {2 Printing function} *)
+ let print_binary_list (single : 'a -> string)
+ (unit : string)
+ (binary : string -> string -> string) (l : 'a list) =
+ let rec aux l =
+ match l with
+ [] -> unit
+ | [t] -> single t
+ | t::q ->
+ let r = aux q in
+ Printf.sprintf "%s" (binary (single t) r)
+ in
+ aux l
+
+ let print_symbol s =
+ match s with
+ | s, None -> Printf.sprintf "%i" s
+ | s , Some u -> Printf.sprintf "%i(unit %i)" s u
+
+ let sprint_ac (single : 'a -> string) (l : 'a mset) =
+ (print_binary_list
+ (fun (x,t) ->
+ if t = 1
+ then single x
+ else Printf.sprintf "%i*%s" t (single x)
+ )
+ "0"
+ (fun x y -> x ^ " , " ^ y)
+ l
+ )
+
+ let print_symbol single s l =
+ match l with
+ [] -> Printf.sprintf "%i" s
+ | _ ->
+ Printf.sprintf "%i(%s)"
+ s
+ (print_binary_list single "" (fun x y -> x ^ "," ^ y) l)
+
+
+ let print_ac single s l =
+ Printf.sprintf "[%s:AC]{%s}"
+ (string_of_int s )
+ (sprint_ac
+ single
+ l
+ )
+
+ let print_a single s l =
+ Printf.sprintf "[%s:A]{%s}"
+ (string_of_int s)
+ (print_binary_list single "1" (fun x y -> x ^ " , " ^ y) l)
+
+ let rec sprint_nf_term = function
+ | TSym (s,l) -> print_symbol sprint_nf_term s l
+ | TAC (s,l) -> print_ac sprint_nf_term s l
+ | TA (s,l) -> print_a sprint_nf_term s l
+ | TVar v -> Printf.sprintf "x%i" v
+ | TUnit s -> Printf.sprintf "unit%i" s
+
+
+
+
+ (** {2 Conversion functions} *)
+
+ (* rebuilds a tree out of a list, under the assumption that the list is not empty *)
+ let rec binary_of_list f comb l =
+ let l = List.rev l in
+ let rec aux = function
+ | [] -> assert false
+ | [t] -> f t
+ | t::q -> comb (aux q) (f t)
+ in
+ aux l
+
+ let term_of_t units : t -> nf_term =
+ let rec term_of_t = function
+ | Dot (s,l,r) ->
+ let l = term_of_t l in
+ let r = term_of_t r in
+ mk_TA units s [l;r]
+ | Plus (s,l,r) ->
+ let l = term_of_t l in
+ let r = term_of_t r in
+ mk_TAC units s [l,1;r,1]
+ | Unit x ->
+ mk_TUnit x
+ | Sym (s,t) ->
+ let t = Array.to_list t in
+ let t = List.map term_of_t t in
+ mk_TSym s t
+ | Var i ->
+ mk_TVar ( i)
+ in
+ term_of_t
+
+ let rec t_of_term : nf_term -> t = function
+ | TAC (s,l) ->
+ (binary_of_list
+ t_of_term
+ (fun l r -> Plus ( s,l,r))
+ (linear l)
+ )
+ | TA (s,l) ->
+ (binary_of_list
+ t_of_term
+ (fun l r -> Dot ( s,l,r))
+ l
+ )
+ | TSym (s,l) ->
+ let v = Array.of_list l in
+ let v = Array.map (t_of_term) v in
+ Sym ( s,v)
+ | TVar x -> Var x
+ | TUnit s -> Unit s
+
+
+ let equal_aac units x y =
+ nf_equal (term_of_t units x) (term_of_t units y)
+ end
+
+ (** Terms environments defined as association lists from variables to
+ terms in normal form {! Terms.nf_term} *)
+ module Subst : sig
+ type t
+
+ val find : t -> var -> Terms.nf_term option
+ val add : t -> var -> Terms.nf_term -> t
+ val empty : t
+ val instantiate : t -> Terms.t -> Terms.t
+ val sprint : t -> string
+ val to_list : t -> (var*Terms.t) list
+ end
+ =
+ struct
+ open Terms
+
+ (** Terms environments, with nf_terms, to avoid costly conversions
+ of {!Terms.nf_terms} to {!Terms.t}, that will be mostly discarded*)
+ type t = (var * nf_term) list
+
+ let find : t -> var -> nf_term option = fun t x ->
+ try Some (List.assoc x t) with | _ -> None
+ let add t x v = (x,v) :: t
+ let empty = []
+
+ let sprint (l : t) =
+ match l with
+ | [] -> Printf.sprintf "Empty environment\n"
+ | _ ->
+ let s = List.fold_left
+ (fun acc (x,y) ->
+ Printf.sprintf "%sX%i -> %s\n"
+ acc
+ x
+ (sprint_nf_term y)
+ )
+ ""
+ (List.rev l) in
+ Printf.sprintf "%s\n%!" s
+
+
+
+ (** [instantiate] is an homomorphism except for the variables*)
+ let instantiate (t: t) (x:Terms.t) : Terms.t =
+ let rec aux = function
+ | Unit _ as x -> x
+ | Sym (s,t) -> Sym (s,Array.map aux t)
+ | Plus (s,l,r) -> Plus (s, aux l, aux r)
+ | Dot (s,l,r) -> Dot (s, aux l, aux r)
+ | Var i ->
+ begin match find t i with
+ | None -> Util.error "aac_tactics: instantiate failure"
+ | Some x -> t_of_term x
+ end
+ in aux x
+
+ let to_list t = List.map (fun (x,y) -> x,Terms.t_of_term y) t
+ end
+
+ (******************)
+ (* MATCHING UTILS *)
+ (******************)
+
+ open Terms
+
+ (** Since most of the folowing functions require an extra parameter,
+ the units, we package them in a module. This functor will then be
+ applied to a module containing the units, in the exported
+ functions. *)
+ module M (X : sig
+ val units : units
+ val is_ac : (symbol * bool) list
+ val strict : bool (* variables cannot be instantiated with units *)
+ end) = struct
+
+ open X
+
+ let nullifiable p =
+ let nullable = not strict in
+ let has_unit s =
+ try let _ = get_unit units s in true
+ with NoUnit -> false
+ in
+ let rec aux = function
+ | TA (s,l) -> has_unit s && List.for_all (aux) l
+ | TAC (s,l) -> has_unit s && List.for_all (fun (x,n) -> aux x) l
+ | TSym _ -> false
+ | TVar _ -> nullable
+ | TUnit _ -> true
+ in aux p
+
+ let print_units ()=
+ List.iter (fun (op,unit) -> Printf.printf "%i %i\t" op unit) units;
+ Printf.printf "\n%!"
+
+ let mk_TAC s l = mk_TAC units s l
+ let mk_TA s l = mk_TA units s l
+ let mk_TAC' s l =
+ try return( mk_TAC s l)
+ with _ -> fail ()
+ let mk_TA' s l =
+ try return( mk_TA s l)
+ with _ -> fail ()
+
+ (** First, we need to be able to perform non-deterministic choice of
+ term splitting to satisfy a pattern. Indeed, we want to show that:
+ (x+a*b)*c <= a*b*c
+ *)
+ let a_nondet_split_raw t : ('a list * 'a list) m =
+ let rec aux l l' =
+ match l' with
+ | [] ->
+ return ( l,[])
+ | t::q ->
+ return ( l,l' )
+ >>| aux (l @ [t]) q
+ in
+ aux [] t
+
+ (** Same as the previous [a_nondet_split], but split the list in 3
+ parts *)
+ let a_nondet_middle t : ('a list * 'a list * 'a list) m =
+ a_nondet_split_raw t >>
+ (fun (left, right) ->
+ a_nondet_split_raw left >>
+ (fun (left, middle) -> return (left, middle, right))
+ )
+
+ (** Non deterministic splits of ac lists *)
+ let dispatch f n =
+ let rec aux k =
+ if k = 0 then return (f n 0)
+ else return (f (n-k) k) >>| aux (k-1)
+ in
+ aux (n )
+
+ let add_with_arith x ar l =
+ if ar = 0 then l else (x,ar) ::l
+
+ let ac_nondet_split_raw (l : 'a mset) : ('a mset * 'a mset) m =
+ let rec aux = function
+ | [] -> return ([],[])
+ | (t,tar)::q ->
+ aux q
+ >>
+ (fun (left,right) ->
+ dispatch (fun arl arr ->
+ add_with_arith t arl left,
+ add_with_arith t arr right
+ )
+ tar
+ )
+ in
+ aux l
+
+ let a_nondet_split current t : (nf_term * nf_term list) m=
+ a_nondet_split_raw t
+ >>
+ (fun (l,r) ->
+ if strict && (l=[])
+ then fail()
+ else
+ mk_TA' current l >>
+ fun t -> return (t, r)
+ )
+
+ let ac_nondet_split current t : (nf_term * nf_term mset) m=
+ ac_nondet_split_raw t
+ >>
+ (fun (l,r) ->
+ if strict && (l=[])
+ then fail()
+ else
+ mk_TAC' current l >>
+ fun t -> return (t, r)
+ )
+
+
+ (** Try to affect the variable [x] to each left factor of [t]*)
+ let var_a_nondet_split env current x t =
+ a_nondet_split current t
+ >>
+ (fun (t,res) ->
+ return ((Subst.add env x t), res)
+ )
+
+ (** Try to affect variable [x] to _each_ subset of t. *)
+ let var_ac_nondet_split (current: symbol) env (x : var) (t : nf_term mset) : (Subst.t * (nf_term mset)) m =
+ ac_nondet_split current t
+ >>
+ (fun (t,res) ->
+ return ((Subst.add env x t), res)
+ )
+
+ (** See the term t as a given AC symbol. Unwrap the first constructor
+ if necessary *)
+ let get_AC (s : symbol) (t : nf_term) : (nf_term *int) list =
+ match t with
+ | TAC (s',l) when s' = s -> l
+ | TUnit u when is_unit units s u -> []
+ | _ -> [t,1]
+
+ (** See the term t as a given A symbol. Unwrap the first constructor
+ if necessary *)
+ let get_A (s : symbol) (t : nf_term) : nf_term list =
+ match t with
+ | TA (s',l) when s' = s -> l
+ | TUnit u when is_unit units s u -> []
+ | _ -> [t]
+
+ (** See the term [t] as an symbol [s]. Fail if it is not such
+ symbol. *)
+ let get_Sym s t =
+ match t with
+ | TSym (s',l) when s' = s -> return l
+ | _ -> fail ()
+
+ (*************)
+ (* A Removal *)
+ (*************)
+
+ (** We remove the left factor v in a term list. This function runs
+ linearly with respect to the size of the first pattern symbol *)
+
+ let left_factor current (v : nf_term) (t : nf_term list) =
+ let rec aux a b =
+ match a,b with
+ | t::q , t' :: q' when nf_equal t t' -> aux q q'
+ | [], q -> return q
+ | _, _ -> fail ()
+ in
+ match v with
+ | TA (s,l) when s = current -> aux l t
+ | TUnit u when is_unit units current u -> return t
+ | _ ->
+ begin match t with
+ | [] -> fail ()
+ | t::q ->
+ if nf_equal v t
+ then return q
+ else fail ()
+ end
+
+
+ (**************)
+ (* AC Removal *)
+ (**************)
+
+ (** {!pick_sym} gather all elements of a list that satisfies a
+ predicate, and combine them with the residual of the list. That
+ is, each element of the residual contains exactly one element less
+ than the original term.
+
+ TODO : This function not as efficient as it could be, using a
+ proper data-structure *)
+
+ let pick_sym (s : symbol) (t : nf_term mset ) =
+ let rec aux front back =
+ match back with
+ | [] -> fail ()
+ | (t,tar)::q ->
+ begin match t with
+ | TSym (s',v') when s = s' ->
+ let back =
+ if tar > 1
+ then (t,tar -1) ::q
+ else q
+ in
+ return (v' , List.rev_append front back )
+ >>| aux ((t,tar)::front) q
+ | _ -> aux ((t,tar)::front) q
+ end
+ in
+ aux [] t
+
+
+
+ (** We have to check if we are trying to remove a unit from a term. Could also be located in Terms*)
+ let is_unit_AC s t =
+ try nf_equal t (mk_TAC s [])
+ with | NoUnit -> false
+
+ let is_same_AC s t : nf_term mset option=
+ match t with
+ TAC (s',l) when s = s' -> Some l
+ | _ -> None
+
+ (** We want to remove the term [v] from the term list [t] under an AC
+ symbol *)
+ let single_AC_factor (s : symbol) (v : nf_term) v_ar (t : nf_term mset) : (nf_term mset) m =
+ let rec aux front back =
+ match back with
+ | [] -> fail ()
+ | (t,tar)::q ->
+ begin
+ if nf_equal v t
+ then
+ match () with
+ | _ when tar < v_ar -> fail ()
+ | _ when tar = v_ar -> return (List.rev_append front q)
+ | _ -> return (List.rev_append front ((t,tar-v_ar)::q))
+ else
+ aux ((t,tar) :: front) q
+ end
+ in
+ if is_unit_AC s v
+ then
+ return t
+ else
+ aux [] t
+
+ (* Remove a constant from a mset. If this constant is also a mset for
+ the same symbol, we remove every elements, one at a time (and we do
+ care of the arities) *)
+ let factor_AC (s : symbol) (v: nf_term) (t : nf_term mset) : ( nf_term mset ) m =
+ match is_same_AC s v with
+ | None -> single_AC_factor s v 1 t
+ | Some l ->
+ (* We are trying to remove an AC factor *)
+ List.fold_left (fun acc (v,v_ar) ->
+ acc >> (single_AC_factor s v v_ar)
+ )
+ (return t)
+ l
+
+
+(** [tri_fold f l acc] folds on the list [l] and give to f the
+ beginning of the list in reverse order, the considered element, and
+ the last part of the list
+
+ as an exemple, on the list [1;2;3;4], we get the trace
+ f () [] 1 [2; 3; 4]
+ f () [1] 2 [3; 4]
+ f () [2;1] 3 [ 4]
+ f () [3;2;1] 4 []
+
+ it is the duty of the user to reverse the front if needed
+*)
+
+let tri_fold f (l : 'a list) (acc : 'b)= match l with
+ [] -> acc
+ | _ ->
+ let _,_,acc = List.fold_left (fun acc (t : 'a) ->
+ let l,r,acc = acc in
+ let r = List.tl r in
+ t::l,r,f acc l t r
+ ) ([], l,acc) l
+ in acc
+
+(* let test = tri_fold (fun acc l t r -> (l,t,r) :: acc) [1;2;3;4] [] *)
+
+
+
+ (*****************************)
+ (* ENUMERATION DES POSITIONS *)
+ (*****************************)
+
+
+(** The pattern is a unit: we need to try to make it appear at each
+ position. We do not need to go further with a real matching, since
+ the match should be trivial. Hence, we proceed recursively to
+ enumerate all the positions. *)
+
+module Positions = struct
+
+
+ let ac (l: 'a mset) : ('a mset * 'a )m =
+ let rec aux = function
+ | [] -> assert false
+ | [t,1] -> return ([],t)
+ | [t,tar] -> return ([t,tar -1],t)
+ | (t,tar) as h :: q ->
+ (aux q >> (fun (c,x) -> return (h :: c,x)))
+ >>| (if tar > 1 then return ((t,tar-1) :: q,t) else return (q,t))
+ in
+ aux l
+
+ let ac' current (l: 'a mset) : ('a mset * 'a )m =
+ ac_nondet_split_raw l >>
+ (fun (l,r) ->
+ if l = [] || r = []
+ then fail ()
+ else
+ mk_TAC' current r >>
+ fun t -> return (l, t)
+ )
+
+ let a (l : 'a list) : ('a list * 'a * 'a list) m =
+ let rec aux left right : ('a list * 'a * 'a list) m =
+ match right with
+ | [] -> assert false
+ | [t] -> return (left,t,[])
+ | t::q ->
+ aux (left@[t]) q
+ >>| return (left,t,q)
+ in
+ aux [] l
+end
+
+let build_ac (current : symbol) (context : nf_term mset) (p : t) : t m=
+ if context = []
+ then return p
+ else
+ mk_TAC' current context >>
+ fun t -> return (Plus (current,t_of_term t,p))
+
+let build_a (current : symbol)
+ (left : nf_term list) (right : nf_term list) (p : t) : t m=
+ let right_pat p =
+ if right = []
+ then return p
+ else
+ mk_TA' current right >>
+ fun t -> return (Dot (current,p,t_of_term t))
+ in
+ let left_pat p=
+ if left = []
+ then return p
+ else
+ mk_TA' current left >>
+ fun t -> return (Dot (current,t_of_term t,p))
+ in
+ right_pat p >> left_pat >> (fun p -> return p)
+
+
+let conts (unit : symbol) (l : symbol list) p : t m =
+ let p = t_of_term p in
+ (* - aller chercher les symboles ac et les symboles a
+ - construire pour chaque
+ * * +
+ / \ / \ / \
+ 1 p p 1 p 1
+ *)
+ let ac,a = List.partition (fun s -> List.assoc s is_ac) l in
+ let acc = fail () in
+ let acc = List.fold_left
+ (fun acc s ->
+ acc >>| return (Plus (s,p,Unit unit))
+ ) acc ac in
+ let acc =
+ List.fold_left
+ (fun acc s ->
+ acc >>| return (Dot (s,p,Unit unit)) >>| return (Dot (s,Unit unit,p))
+ ) acc a
+ in acc
+
+
+(**
+
+ Return the same thing as subterm :
+ - The size of the context
+ - The context
+ - A collection of substitutions (here == return Subst.empty)
+*)
+let unit_subterm (t : nf_term) (unit : symbol) : (int * t * Subst.t m) m =
+ let symbols = List.fold_left
+ (fun acc (ac,u) -> if u = unit then ac :: acc else acc ) [] units
+ in
+ (* make a unit appear at each strict sub-position of the term*)
+ let rec positions (t : nf_term) : t m =
+ match t with
+ (* with a final unit at the end *)
+ | TAC (s,l) ->
+ let symbols' = List.filter (fun x -> x <> s) symbols in
+ (
+ ac_nondet_split_raw l >>
+ (fun (l,r) -> if l = [] || r = [] then fail () else
+ (
+ match r with
+ | [p,1] ->
+ positions p >>| conts unit symbols' p
+ | _ ->
+ mk_TAC' s r >> conts unit symbols'
+ ) >> build_ac s l ))
+ | TA (s,l) ->
+ let symbols' = List.filter (fun x -> x <> s) symbols in
+ (
+ (* first the other symbols, and then the more simple case of
+ this aprticular symbol *)
+ a_nondet_middle l >>
+ (fun (l,m,r) ->
+ (* meant to break the symmetry *)
+ if (l = [] && r = [])
+ then fail ()
+ else
+ (
+ match m with
+ | [p] ->
+ positions p >>| conts unit symbols' p
+ | _ ->
+ mk_TA' s m >> conts unit symbols'
+ ) >> build_a s l r ))
+ >>|
+ (
+ if List.mem s symbols then
+ begin match l with
+ | [a] -> assert false
+ | [a;b] -> build_a s [a] [b] (Unit unit)
+ | _ ->
+ (* on ne construit que les elements interieurs,
+ d'ou la disymetrie *)
+ (Positions.a l >>
+ (fun (left,p,right) ->
+ if left = [] then fail () else
+ (build_a s left right (Dot (s,Unit unit,t_of_term p)))))
+ end
+ else fail ()
+ )
+
+ | TSym (s,l) ->
+ tri_fold (fun acc l t r ->
+ ((positions t) >>
+ (fun (p) ->
+ let l = List.map t_of_term l in
+ let r = List.map t_of_term r in
+ return (Sym (s, Array.of_list (List.rev_append l (p::r)))) ))
+ >>|
+ (
+ conts unit symbols t >>
+ (fun (p) ->
+ let l = List.map t_of_term l in
+ let r = List.map t_of_term r in
+ return (Sym (s, Array.of_list (List.rev_append l (p::r)))) )
+ )
+ >>| acc
+ ) l (fail())
+ | TVar x -> assert false
+ | TUnit x when x = unit -> return (Unit unit)
+ | TUnit x as t -> conts unit symbols t
+ in
+ (positions t
+ >>|
+ (match t with
+ | TSym _ -> conts unit symbols t
+ | TAC (s,l) -> conts unit symbols t
+ | TA (s,l) -> conts unit symbols t
+ | _ -> fail())
+ )
+ >> fun (p) -> return (Terms.size p,p,return Subst.empty)
+
+
+
+
+(************)
+ (* Matching *)
+ (************)
+
+
+
+(** {!matching} is the generic matching judgement. Each time a
+ non-deterministic split is made, we have to come back to this one.
+
+ {!matchingSym} is used to match two applications that have the
+ same (free) head-symbol.
+
+ {!matchingAC} is used to match two sums (with the subtlety that
+ [x+y] matches [f a] which is a function application or [a*b] which
+ is a product).
+
+ {!matchingA} is used to match two products (with the subtlety that
+ [x*y] matches [f a] which is a function application, or [a+b]
+ which is a sum).
+
+
+*)
+
+let matching =
+ let rec matching env (p : nf_term) (t: nf_term) : Subst.t AAC_search_monad.m=
+ match p with
+ | TAC (s,l) ->
+ let l = linear l in
+ matchingAC env s l (get_AC s t)
+ | TA (s,l) ->
+ matchingA env s l (get_A s t)
+ | TSym (s,l) ->
+ (get_Sym s t)
+ >> (fun t -> matchingSym env l t)
+ | TVar x ->
+ begin match Subst.find env x with
+ | None -> return (Subst.add env x t)
+ | Some v -> if nf_equal v t then return env else fail ()
+ end
+ | TUnit s ->
+ if nf_equal p t then return env else fail ()
+ and
+ matchingAC (env : Subst.t) (current : symbol) (l : nf_term list) (t : (nf_term *int) list) =
+ match l with
+ | TSym (s,v):: q ->
+ pick_sym s t
+ >> (fun (v',t') ->
+ matchingSym env v v'
+ >> (fun env -> matchingAC env current q t'))
+
+ | TAC (s,v)::q when s = current ->
+ assert false
+ | TVar x:: q -> (* This is an optimization *)
+ begin match Subst.find env x with
+ | None ->
+ (var_ac_nondet_split current env x t)
+ >> (fun (env,residual) -> matchingAC env current q residual)
+ | Some v ->
+ (factor_AC current v t)
+ >> (fun residual -> matchingAC env current q residual)
+ end
+ | TUnit s as v :: q -> (* this is an optimization *)
+ (factor_AC current v t) >>
+ (fun residual -> matchingAC env current q residual)
+ | h :: q ->(* PAC =/= curent or PA or unit for this symbol*)
+ (ac_nondet_split current t)
+ >>
+ (fun (t,right) ->
+ matching env h t
+ >>
+ (
+ fun env ->
+ matchingAC env current q right
+ )
+ )
+ | [] -> if t = [] then return env else fail ()
+ and
+ matchingA (env : Subst.t) (current : symbol) (l : nf_term list) (t : nf_term list) =
+ match l with
+ | TSym (s,v) :: l ->
+ begin match t with
+ | TSym (s',v') :: r when s = s' ->
+ (matchingSym env v v')
+ >> (fun env -> matchingA env current l r)
+ | _ -> fail ()
+ end
+ | TA (s,v) :: l when s = current ->
+ assert false
+ | TVar x :: l ->
+ begin match Subst.find env x with
+ | None ->
+ debug (Printf.sprintf "var %i (%s)" x
+ (let b = Buffer.create 21 in List.iter (fun t -> Buffer.add_string b ( Terms.sprint_nf_term t)) t; Buffer.contents b ));
+ var_a_nondet_split env current x t
+ >> (fun (env,residual)-> debug (Printf.sprintf "pl %i %i" x(List.length residual)); matchingA env current l residual)
+ | Some v ->
+ (left_factor current v t)
+ >> (fun residual -> matchingA env current l residual)
+ end
+ | TUnit s as v :: q -> (* this is an optimization *)
+ (left_factor current v t) >>
+ (fun residual -> matchingA env current q residual)
+ | h :: l ->
+ a_nondet_split current t
+ >> (fun (t,r) ->
+ matching env h t
+ >> (fun env -> matchingA env current l r)
+ )
+ | [] -> if t = [] then return env else fail ()
+ and
+ matchingSym (env : Subst.t) (l : nf_term list) (t : nf_term list) =
+ List.fold_left2
+ (fun acc p t -> acc >> (fun env -> matching env p t))
+ (return env)
+ l
+ t
+
+ in
+ fun env l r ->
+ let _ = debug (Printf.sprintf "pattern :%s\nterm :%s\n%!" (Terms.sprint_nf_term l) (Terms.sprint_nf_term r)) in
+ let m = matching env l r in
+ let _ = debug (Printf.sprintf "count %i" (AAC_search_monad.count m)) in
+ m
+
+
+
+(***********)
+(* Subterm *)
+(***********)
+
+
+
+(** [subterm] solves a sub-term pattern matching.
+
+ This function is more high-level than {!matcher}, thus takes {!t}
+ as arguments rather than terms in normal form {!nf_term}.
+
+ We use three mutually recursive functions {!subterm},
+ {!subterm_AC}, {!subterm_A} to find the matching subterm, making
+ non-deterministic choices to split the term into a context and an
+ intersting sub-term. Intuitively, the only case in which we have to
+ go in depth is when we are left with a sub-term that is atomic.
+
+ Indeed, rewriting [H: b = c |- a+b+a = a+a+c], we do not want to
+ find recursively the sub-terms of [a+b] and [b+a], since they will
+ overlap with the sub-terms of [a+b+a].
+
+ We rebuild the context on the fly, leaving the variables in the
+ pattern uninstantiated. We do so in order to allow interaction
+ with the user, to choose the env.
+
+ Strange patterms like x*y*x can be instanciated by nothing, inside
+ a product. Therefore, we need to check that all the term is not
+ going into the context. With proper support for interaction with
+ the user, we should lift these tests. However, at the moment, they
+ serve as heuristics to return "interesting" matchings
+*)
+
+let return_non_empty raw_p m =
+ if is_empty m
+ then
+ fail ()
+ else
+ return (raw_p ,m)
+
+let subterm (raw_p:t) (raw_t:t): (int * t * Subst.t m) m=
+ let _ = debug (String.make 40 '=') in
+ let p = term_of_t units raw_p in
+ let t = term_of_t units raw_t in
+ let _ =
+ if nullifiable p
+ then
+ Pp.msg_warning
+ (Pp.str
+ "[aac_tactics] This pattern might be nullifiable, some solutions can be missing");
+ in
+
+ let _ = debug (Printf.sprintf "%s" (Terms.sprint_nf_term p)) in
+ let _ = debug (Printf.sprintf "%s" (Terms.sprint_nf_term t)) in
+ let filter_right current right (p,m) =
+ if right = []
+ then return (p,m)
+ else
+ mk_TAC' current right >>
+ fun t -> return (Plus (current,p,t_of_term t),m)
+ in
+ let filter_middle current left right (p,m) =
+ let right_pat p =
+ if right = []
+ then return p
+ else
+ mk_TA' current right >>
+ fun t -> return (Dot (current,p,t_of_term t))
+ in
+ let left_pat p=
+ if left = []
+ then return p
+ else
+ mk_TA' current left >>
+ fun t -> return (Dot (current,t_of_term t,p))
+ in
+ right_pat p >> left_pat >> (fun p -> return (p,m))
+ in
+ let rec subterm (t:nf_term) : (t * Subst.t m) m=
+ match t with
+ | TAC (s,l) ->
+ ((ac_nondet_split_raw l) >>
+ (fun (left,right) ->
+ (subterm_AC s left) >> (filter_right s right)
+ ))
+ | TA (s,l) ->
+ (a_nondet_middle l) >>
+ (fun (left, middle, right) ->
+ (subterm_A s middle) >>
+ (filter_middle s left right)
+ )
+ | TSym (s, l) ->
+ let init =
+ return_non_empty raw_p (matching Subst.empty p t)
+ in
+ tri_fold (fun acc l t r ->
+ ((subterm t) >>
+ (fun (p,m) ->
+ let l = List.map t_of_term l in
+ let r = List.map t_of_term r in
+ let p = Sym (s, Array.of_list (List.rev_append l (p::r))) in
+ return (p,m)
+ )) >>| acc
+ ) l init
+ | TVar x -> assert false
+ | TUnit s -> fail ()
+
+ and subterm_AC s tl =
+ match tl with
+ [x,1] -> subterm x
+ | _ ->
+ mk_TAC' s tl >> fun t ->
+ return_non_empty raw_p (matching Subst.empty p t)
+ and subterm_A s tl =
+ match tl with
+ [x] -> subterm x
+ | _ ->
+ mk_TA' s tl >>
+ fun t ->
+ return_non_empty raw_p (matching Subst.empty p t)
+ in
+ match p with
+ | TUnit x ->
+ unit_subterm t x
+ | _ -> (subterm t >> fun (p,m) -> return (Terms.size p,p,m))
+
+ end
+
+(* The functions we export, handlers for the previous ones. Some debug
+ information also *)
+let subterm ?(strict = false) units raw t =
+ let module M = M (struct
+ let is_ac = units.is_ac
+ let units = units.unit_for
+ let strict = strict
+ end) in
+ let sols = time (M.subterm raw) t "%fs spent in subterm (including matching)\n" in
+ debug
+ (Printf.sprintf "%i possible solution(s)\n"
+ (AAC_search_monad.fold (fun (_,_,envm) acc -> count envm + acc) sols 0));
+ sols
+
+
+let matcher ?(strict = false) units p t =
+ let module M = M (struct
+ let is_ac = units.is_ac
+ let units = units.unit_for
+ let strict = false
+ end) in
+ let units = units.unit_for in
+ let sols = time
+ (fun (p,t) ->
+ let p = (Terms.term_of_t units p) in
+ let t = (Terms.term_of_t units t) in
+ M.matching Subst.empty p t) (p,t)
+ "%fs spent in the matcher\n"
+ in
+ debug (Printf.sprintf "%i solutions\n" (count sols));
+ sols
+
diff --git a/matcher.mli b/AAC_matcher.mli
index 3e20e73..896a2a9 100644
--- a/matcher.mli
+++ b/AAC_matcher.mli
@@ -7,27 +7,42 @@
(***************************************************************************)
(** Standalone module containing the algorithm for matching modulo
- associativity and associativity and commutativity (AAC).
+ associativity and associativity and commutativity
+ (AAC). Additionnaly, some A or AC operators can have units (U).
- This module could be reused ouside of the Coq plugin.
+ This module could be reused outside of the Coq plugin.
- Matching modulo AAC a pattern [p] against a term [t] boils down to
- finding a substitution [env] such that the pattern [p] instantiated
- with [env] is equal to [t] modulo AAC.
+ Matching a pattern [p] against a term [t] modulo AACU boils down
+ to finding a substitution [env] such that the pattern [p]
+ instantiated with [env] is equal to [t] modulo AACU.
We proceed by structural decomposition of the pattern, trying all
- possible non-deterministic split of the subject, when needed. The
+ possible non-deterministic splittings of the subject, when needed. The
function {!matcher} is limited to top-level matching, that is, the
subject must make a perfect match against the pattern ([x+x] does
not match [a+a+b] ).
-
- We use a search monad {!Search} to perform non-deterministic
+
+ We use a search monad {!AAC_search_monad} to perform non-deterministic
choices in an almost transparent way.
We also provide a function {!subterm} for finding a match that is
- a subterm of the subject modulo AAC. In particular, this function
+ a subterm of the subject modulo AACU. In particular, this function
gives a solution to the aforementioned case ([x+x] against
[a+b+a]).
+
+ On a slightly more involved level :
+ - it must be noted that we allow several AC/A operators to share
+ the same units, but that a given AC/A operator can have at most
+ one unit.
+
+ - if the pattern does not contain "hard" symbols (like constants,
+ function symbols, AC or A symbols without units), there can be
+ infinitely many subterms such that the pattern matches: it is
+ possible to build "subterms" modulo AAC and U that make the size
+ of the term increase (by making neutral elements appear in a
+ layered fashion). Hence, in this case, a warning is issued, and
+ some solutions are omitted.
+
*)
(** {2 Utility functions} *)
@@ -35,34 +50,19 @@
type symbol = int
type var = int
-(** The {!Search} module contains a search monad that allows to
- express our non-deterministic and back-tracking algorithm in a
- legible maner.
+(** Relationship between units and operators. This is a sparse
+ representation of a matrix of couples [(op,unit)] where [op] is
+ the index of the operation, and [unit] the index of the relevant
+ unit. We make the assumption that any operation has 0 or 1 unit,
+ and that operations can share a unit). *)
+
+type units =(symbol * symbol) list (* from AC/A symbols to the unit *)
+type ext_units =
+ {
+ unit_for : units; (* from AC/A symbols to the unit *)
+ is_ac : (symbol * bool) list
+ }
- @see <http://spivey.oriel.ox.ac.uk/mike/search-jfp.pdf> the
- inspiration of this module
-*)
-module Search :
-sig
- (** A data type that represent a collection of ['a] *)
- type 'a m
- (** bind and return *)
- val ( >> ) : 'a m -> ('a -> 'b m) -> 'b m
- val return : 'a -> 'a m
- (** non-deterministic choice *)
- val ( >>| ) : 'a m -> 'a m -> 'a m
- (** failure *)
- val fail : unit -> 'a m
- (** folding through the collection *)
- val fold : ('a -> 'b -> 'b) -> 'a m -> 'b -> 'b
- (** derived facilities *)
- val sprint : ('a -> string) -> 'a m -> string
- val count : 'a m -> int
- val choose : 'a m -> 'a option
- val to_list : 'a m -> 'a list
- val sort : ('a -> 'a -> int) -> 'a m -> 'a m
- val is_empty: 'a m -> bool
-end
(** The arguments of sums (or AC operators) are represented using finite multisets.
(Typically, [a+b+a] corresponds to [2.a+b], i.e. [Sum[a,2;b,1]]) *)
@@ -73,11 +73,11 @@ val linear : 'a mset -> 'a list
(** Representations of expressions
- The module {!Terms} defines two different types for expressions.
+ The module {!Terms} defines two different types for expressions.
- a public type {!Terms.t} that represents abstract syntax trees
of expressions with binary associative and commutative operators
- a private type {!Terms.nf_term}, corresponding to a canonical
- representation for the above terms modulo AAC. The construction
+ representation for the above terms modulo AACU. The construction
functions on this type ensure that these terms are in normal form
(that is, no sum can appear as a subterm of the same sum, no
trailing units, lists corresponding to multisets are sorted,
@@ -97,42 +97,41 @@ sig
by the following value
[Sym(0,[| Dot(1, Sym(2,[||]), Sym(3,[||]));
Dot(4, Sym(2,[||]), Sym(5,[|Sym(3,[||])|])) |])]
- where the implicit symbol environment associates
- [f] to [0], [(^)] to [1], [a] to [2], [b] to [3], [( * )] to [4], and [g] to [5],
+ where the implicit symbol environment associates
+ [f] to [0], [(^)] to [1], [a] to [2], [b] to [3], [( * )] to [4], and [g] to [5],
- Accordingly, the following value, that contains "variables"
- [Sym(0,[| Dot(1, Var x, Dot(4,[||]));
- Dot(4, Var x, Sym(5,[|Sym(3,[||])|])) |])]
- represents the pattern [forall x, f (x^1) (x*g b)], where [1] is the
- unit associated with [( * )].
- *)
+ Accordingly, the following value, that contains "variables"
+ [Sym(0,[| Dot(1, Var x, Unit (1); Dot(4, Var x,
+ Sym(5,[|Sym(3,[||])|])) |])] represents the pattern [forall x, f
+ (x^1) (x*g b)]. The relationship between [1] and [( * )] is only
+ mentionned in the units table. *)
type t =
Dot of (symbol * t * t)
- | One of symbol
| Plus of (symbol * t * t)
- | Zero of symbol
| Sym of (symbol * t array)
| Var of var
+ | Unit of symbol
- (** Test for equality of terms modulo AAC (relies on the following
- canonical representation of terms) *)
- val equal_aac : t -> t -> bool
+ (** Test for equality of terms modulo AACU (relies on the following
+ canonical representation of terms). Note than two different
+ units of a same operator are not considered equal. *)
+ val equal_aac : units -> t -> t -> bool
(** {2 Normalised terms (canonical representation) }
-
+
A term in normal form is the canonical representative of the
- equivalence class of all the terms that are equal modulo AAC
+ equivalence class of all the terms that are equal modulo AACU.
This representation is only used internally; it is exported here
for the sake of completeness *)
- type nf_term
+ type nf_term
(** {3 Comparisons} *)
val nf_term_compare : nf_term -> nf_term -> int
- val nf_equal : nf_term -> nf_term -> bool
+ val nf_equal : nf_term -> nf_term -> bool
(** {3 Printing function} *)
@@ -140,10 +139,10 @@ sig
(** {3 Conversion functions} *)
- (** we have the following property: [a] and [b] are equal modulo AAC
+ (** we have the following property: [a] and [b] are equal modulo AACU
iif [nf_equal (term_of_t a) (term_of_t b) = true] *)
- val term_of_t : t -> nf_term
- val t_of_term : nf_term -> t
+ val term_of_t : units -> t -> nf_term
+ val t_of_term : nf_term -> t
end
@@ -153,11 +152,11 @@ end
The module {!Subst} contains infrastructure to deal with
substitutions, i.e., functions from variables to terms. Only a
restricted subsets of these functions need to be exported.
-
+
As expected, a particular substitution can be used to
instantiate a pattern.
*)
-module Subst :
+module Subst :
sig
type t
val sprint : t -> string
@@ -176,17 +175,15 @@ end
while x+x+y does not match a+b+c, since this would require to
assign 1 to x).
*)
-val matcher : ?strict:bool -> Terms.t -> Terms.t -> Subst.t Search.m
+val matcher : ?strict:bool -> ext_units -> Terms.t -> Terms.t -> Subst.t AAC_search_monad.m
(** [subterm p t] computes a set of solutions to the given
subterm-matching problem.
-
- @return a collection of possible solutions (each with the
+
+ Return a collection of possible solutions (each with the
associated depth, the context, and the solutions of the matching
problem). The context is actually a {!Terms.t} where the variables
are yet to be instantiated by one of the associated substitutions
*)
-val subterm : ?strict:bool -> Terms.t -> Terms.t -> (int * Terms.t * Subst.t Search.m) Search.m
+val subterm : ?strict:bool -> ext_units -> Terms.t -> Terms.t -> (int * Terms.t * Subst.t AAC_search_monad.m) AAC_search_monad.m
-(** pretty printing of the solutions *)
-val pp_all : (Terms.t -> Pp.std_ppcmds) -> (int * Terms.t * Subst.t Search.m) Search.m -> Pp.std_ppcmds
diff --git a/AAC_print.ml b/AAC_print.ml
new file mode 100644
index 0000000..512efa3
--- /dev/null
+++ b/AAC_print.ml
@@ -0,0 +1,100 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+(* A very basic way to interact with the envs, to choose a possible
+ solution *)
+open Pp
+open AAC_matcher
+type named_env = (Names.name * Terms.t) list
+
+
+
+(** {pp_env} prints a substitution mapping names to terms, using the
+provided printer *)
+let pp_env pt : named_env -> Pp.std_ppcmds = fun env ->
+ List.fold_left
+ (fun acc (v,t) ->
+ begin match v with
+ | Names.Name s -> str (Printf.sprintf "%s: " (Names.string_of_id s))
+ | Names.Anonymous -> str ("_")
+ end
+ ++ pt t ++ str "; " ++ acc
+ )
+ (str "")
+ env
+
+(** {pp_envm} prints a collection of possible environments, and number
+them. This number must remain compatible with the parameters given to
+{!aac_rewrite} *)
+let pp_envm pt : named_env AAC_search_monad.m -> Pp.std_ppcmds = fun m ->
+ let _,s =
+ AAC_search_monad.fold
+ (fun env (n,acc) ->
+ n+1, h 0 (str (Printf.sprintf "%i:\t[" n) ++pp_env pt env ++ str "]") ++ fnl () :: acc
+ ) m (0,[])
+ in
+ List.fold_left (fun acc s -> s ++ acc) (str "") (s)
+
+let trivial_substitution envm =
+ match AAC_search_monad.choose envm with
+ | None -> true (* assert false *)
+ | Some l -> l=[]
+
+(** {pp_all} prints a collection of possible contexts and related
+possibles substitutions, giving a number to each. This number must
+remain compatible with the parameters of {!aac_rewrite} *)
+let pp_all pt : (int * Terms.t * named_env AAC_search_monad.m) AAC_search_monad.m -> Pp.std_ppcmds = fun m ->
+ let _,s = AAC_search_monad.fold
+ (fun (size,context,envm) (n,acc) ->
+ let s = str (Printf.sprintf "occurence %i: transitivity through " n) in
+ let s = s ++ pt context ++ str "\n" in
+ let s = if trivial_substitution envm then s else
+ s ++ str (Printf.sprintf "%i possible(s) substitution(s)" (AAC_search_monad.count envm) )
+ ++ fnl () ++ pp_envm pt envm
+ in
+ n+1, s::acc
+ ) m (0,[]) in
+ List.fold_left (fun acc s -> s ++ str "\n" ++ acc) (str "") (s)
+
+(** The main printing function. {!print} uses the debruijn_env the
+rename the variables, and rebuilds raw Coq terms (for the context, and
+the terms in the environment). In order to do so, it requires the
+information gathered by the {!AAC_theory.Trans} module.*)
+let print rlt ir m (context : Term.rel_context) goal =
+ if AAC_search_monad.count m = 0
+ then
+ (
+ Tacticals.tclFAIL 0 (Pp.str "No subterm modulo AC") goal
+ )
+ else
+ let _ = Pp.msgnl (Pp.str "All solutions:") in
+ let m = AAC_search_monad.(>>) m
+ (fun (i,t,envm) ->
+ let envm = AAC_search_monad.(>>) envm ( fun env ->
+ let l = AAC_matcher.Subst.to_list env in
+ let l = List.sort (fun (n,_) (n',_) -> Pervasives.compare n n') l in
+ let l =
+ List.map (fun (v,t) ->
+ let (name,body,types) = Term.lookup_rel v context in
+ (name,t)
+ ) l
+ in
+ AAC_search_monad.return l
+ )
+ in
+ AAC_search_monad.return (i,t,envm)
+ )
+ in
+ let m = AAC_search_monad.sort (fun (x,_,_) (y,_,_) -> x - y) m in
+ let _ = Pp.msgnl
+ (pp_all
+ (fun t -> Printer.pr_constr (AAC_theory.Trans.raw_constr_of_t ir rlt context t) ) m
+ )
+ in
+ Tacticals.tclIDTAC goal
+
diff --git a/AAC_print.mli b/AAC_print.mli
new file mode 100644
index 0000000..bacd806
--- /dev/null
+++ b/AAC_print.mli
@@ -0,0 +1,23 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+(** Pretty printing functions we use for the aac_instances
+ tactic. *)
+
+
+(** The main printing function. {!print} uses the [Term.rel_context]
+ to rename the variables, and rebuilds raw Coq terms (for the given
+ context, and the terms in the environment). In order to do so, it
+ requires the information gathered by the {!AAC_theory.Trans} module.*)
+val print :
+ AAC_coq.Relation.t ->
+ AAC_theory.Trans.ir ->
+ (int * AAC_matcher.Terms.t * AAC_matcher.Subst.t AAC_search_monad.m) AAC_search_monad.m ->
+ Term.rel_context ->
+ Proof_type.tactic
+
diff --git a/AAC_rewrite.ml b/AAC_rewrite.ml
new file mode 100644
index 0000000..be286f9
--- /dev/null
+++ b/AAC_rewrite.ml
@@ -0,0 +1,524 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+(** aac_rewrite -- rewriting modulo *)
+
+
+module Control = struct
+ let debug = false
+ let printing = false
+ let time = false
+end
+
+module Debug = AAC_helper.Debug (Control)
+open Debug
+
+let time_tac msg tac =
+ if Control.time then AAC_coq.tclTIME msg tac else tac
+
+let tac_or_exn tac exn msg = fun gl ->
+ try tac gl with e ->
+ pr_constr "last goal" (Tacmach.pf_concl gl);
+ exn msg e
+
+(* helper to be used with the previous function: raise a new anomaly
+ except if a another one was previously raised *)
+let push_anomaly msg = function
+ | Util.Anomaly _ as e -> raise e
+ | _ -> AAC_coq.anomaly msg
+
+module M = AAC_matcher
+
+open Term
+open Names
+open Coqlib
+open Proof_type
+
+(** The various kind of relation we can encounter, as a hierarchy *)
+type rew_relation =
+ | Bare of AAC_coq.Relation.t
+ | Transitive of AAC_coq.Transitive.t
+ | Equivalence of AAC_coq.Equivalence.t
+
+(** {!promote try to go higher in the aforementionned hierarchy} *)
+let promote (rlt : AAC_coq.Relation.t) (k : rew_relation -> Proof_type.tactic) =
+ try AAC_coq.Equivalence.cps_from_relation rlt
+ (fun e -> k (Equivalence e))
+ with
+ | Not_found ->
+ begin
+ try AAC_coq.Transitive.cps_from_relation rlt
+ (fun trans -> k (Transitive trans))
+ with
+ |Not_found -> k (Bare rlt)
+ end
+
+
+(*
+ Various situations:
+ p == q |- left == right : rewrite <- ->
+ p <= q |- left <= right : rewrite ->
+ p <= q |- left == right : failure
+ p == q |- left <= right : rewrite <- ->
+
+ Not handled
+ p <= q |- left >= right : failure
+*)
+
+(** aac_lift : the ideal type beyond AAC.v/AAC_lift
+
+ A base relation r, together with an equivalence relation, and the
+ proof that the former lifts to the later. Howver, we have to
+ ensure manually the invariant : r.carrier == e.carrier, and that
+ lift connects the two things *)
+type aac_lift =
+ {
+ r : AAC_coq.Relation.t;
+ e : AAC_coq.Equivalence.t;
+ lift : Term.constr
+ }
+
+type rewinfo =
+ {
+ hypinfo : AAC_coq.Rewrite.hypinfo;
+
+ in_left : bool; (** are we rewriting in the left hand-sie of the goal *)
+ pattern : constr;
+ subject : constr;
+ morph_rlt : AAC_coq.Relation.t; (** the relation we look for in morphism *)
+ eqt : AAC_coq.Equivalence.t; (** the equivalence we use as workbase *)
+ rlt : AAC_coq.Relation.t; (** the relation in the goal *)
+ lifting: aac_lift
+ }
+
+let infer_lifting (rlt: AAC_coq.Relation.t) (k : lift:aac_lift -> Proof_type.tactic) : Proof_type.tactic =
+ AAC_coq.cps_evar_relation rlt.AAC_coq.Relation.carrier (fun e ->
+ let lift_ty =
+ mkApp (Lazy.force AAC_theory.Stubs.lift,
+ [|
+ rlt.AAC_coq.Relation.carrier;
+ rlt.AAC_coq.Relation.r;
+ e
+ |]
+ ) in
+ AAC_coq.cps_resolve_one_typeclass ~error:"Cannot infer a lifting"
+ lift_ty (fun lift goal ->
+ let x = rlt.AAC_coq.Relation.carrier in
+ let r = rlt.AAC_coq.Relation.r in
+ let eq = (AAC_coq.nf_evar goal e) in
+ let equiv = AAC_coq.lapp AAC_theory.Stubs.lift_proj_equivalence [| x;r;eq; lift |] in
+ let lift =
+ {
+ r = rlt;
+ e = AAC_coq.Equivalence.make x eq equiv;
+ lift = lift;
+ }
+ in
+ k ~lift:lift goal
+ ))
+
+(** Builds a rewinfo, once and for all *)
+let dispatch in_left (left,right,rlt) hypinfo (k: rewinfo -> Proof_type.tactic ) : Proof_type.tactic=
+ let l2r = hypinfo.AAC_coq.Rewrite.l2r in
+ infer_lifting rlt
+ (fun ~lift ->
+ let eq = lift.e in
+ k {
+ hypinfo = hypinfo;
+ in_left = in_left;
+ pattern = if l2r then hypinfo.AAC_coq.Rewrite.left else hypinfo.AAC_coq.Rewrite.right;
+ subject = if in_left then left else right;
+ morph_rlt = AAC_coq.Equivalence.to_relation eq;
+ eqt = eq;
+ lifting = lift;
+ rlt = rlt
+ }
+ )
+
+
+
+(** {1 Tactics} *)
+
+
+(** Build the reifiers, the reified terms, and the evaluation fonction *)
+let handle eqt zero envs (t : AAC_matcher.Terms.t) (t' : AAC_matcher.Terms.t) k =
+
+ let (x,r,_) = AAC_coq.Equivalence.split eqt in
+ AAC_theory.Trans.mk_reifier (AAC_coq.Equivalence.to_relation eqt) zero envs
+ (fun (maps, reifier) ->
+ (* fold through a term and reify *)
+ let t = AAC_theory.Trans.reif_constr_of_t reifier t in
+ let t' = AAC_theory.Trans.reif_constr_of_t reifier t' in
+ (* Some letins *)
+ let eval = (mkApp (Lazy.force AAC_theory.Stubs.eval, [|x;r; maps.AAC_theory.Trans.env_sym; maps.AAC_theory.Trans.env_bin; maps.AAC_theory.Trans.env_units|])) in
+
+ AAC_coq.cps_mk_letin "eval" eval (fun eval ->
+ AAC_coq.cps_mk_letin "left" t (fun t ->
+ AAC_coq.cps_mk_letin "right" t' (fun t' ->
+ k maps eval t t'))))
+
+(** [by_aac_reflexivity] is a sub-tactic that closes a sub-goal that
+ is merely a proof of equality of two terms modulo AAC *)
+let by_aac_reflexivity zero
+ eqt envs (t : AAC_matcher.Terms.t) (t' : AAC_matcher.Terms.t) : Proof_type.tactic =
+ handle eqt zero envs t t'
+ (fun maps eval t t' ->
+ let (x,r,e) = AAC_coq.Equivalence.split eqt in
+ let decision_thm = AAC_coq.lapp AAC_theory.Stubs.decide_thm
+ [|x;r;e;
+ maps.AAC_theory.Trans.env_sym;
+ maps.AAC_theory.Trans.env_bin;
+ maps.AAC_theory.Trans.env_units;
+ t;t';
+ |]
+ in
+ (* This convert is required to deal with evars in a proper
+ way *)
+ let convert_to = mkApp (r, [| mkApp (eval,[| t |]); mkApp (eval, [|t'|])|]) in
+ let convert = Tactics.convert_concl convert_to Term.VMcast in
+ let apply_tac = Tactics.apply decision_thm in
+ (Tacticals.tclTHENLIST
+ [
+ convert ;
+ tac_or_exn apply_tac AAC_coq.user_error "unification failure";
+ tac_or_exn (time_tac "vm_norm" (Tactics.normalise_in_concl)) AAC_coq.anomaly "vm_compute failure";
+ Tacticals.tclORELSE Tactics.reflexivity
+ (Tacticals.tclFAIL 0 (Pp.str "Not an equality modulo A/AC"))
+ ])
+ )
+
+let by_aac_normalise zero lift ir t t' =
+ let eqt = lift.e in
+ let rlt = lift.r in
+ handle eqt zero ir t t'
+ (fun maps eval t t' ->
+ let (x,r,e) = AAC_coq.Equivalence.split eqt in
+ let normalise_thm = AAC_coq.lapp AAC_theory.Stubs.lift_normalise_thm
+ [|x;r;e;
+ maps.AAC_theory.Trans.env_sym;
+ maps.AAC_theory.Trans.env_bin;
+ maps.AAC_theory.Trans.env_units;
+ rlt.AAC_coq.Relation.r;
+ lift.lift;
+ t;t';
+ |]
+ in
+ (* This convert is required to deal with evars in a proper
+ way *)
+ let convert_to = mkApp (rlt.AAC_coq.Relation.r, [| mkApp (eval,[| t |]); mkApp (eval, [|t'|])|]) in
+ let convert = Tactics.convert_concl convert_to Term.VMcast in
+ let apply_tac = Tactics.apply normalise_thm in
+ (Tacticals.tclTHENLIST
+ [
+ convert ;
+ apply_tac;
+ ])
+
+ )
+
+(** A handler tactic, that reifies the goal, and infer the liftings,
+ and then call its continuation *)
+let aac_conclude
+ (k : Term.constr -> aac_lift -> AAC_theory.Trans.ir -> AAC_matcher.Terms.t -> AAC_matcher.Terms.t -> Proof_type.tactic) = fun goal ->
+
+ let (equation : Term.types) = Tacmach.pf_concl goal in
+ let envs = AAC_theory.Trans.empty_envs () in
+ let left, right,r =
+ match AAC_coq.match_as_equation goal equation with
+ | None -> AAC_coq.user_error "The goal is not an applied relation"
+ | Some x -> x in
+ try infer_lifting r
+ (fun ~lift goal ->
+ let eq = AAC_coq.Equivalence.to_relation lift.e in
+ let tleft,tright, goal = AAC_theory.Trans.t_of_constr goal eq envs (left,right) in
+ let goal, ir = AAC_theory.Trans.ir_of_envs goal eq envs in
+ let concl = Tacmach.pf_concl goal in
+ let _ = pr_constr "concl "concl in
+ let evar_map = Tacmach.project goal in
+ Tacticals.tclTHEN
+ (Refiner.tclEVARS evar_map)
+ (k left lift ir tleft tright)
+ goal
+ )goal
+ with
+ | Not_found -> AAC_coq.user_error "No lifting from the goal's relation to an equivalence"
+
+open Libnames
+open Tacinterp
+
+let aac_normalise = fun goal ->
+ let ids = Tacmach.pf_ids_of_hyps goal in
+ Tacticals.tclTHENLIST
+ [
+ aac_conclude by_aac_normalise;
+ Tacinterp.interp (
+ <:tactic<
+ intro x;
+ intro y;
+ vm_compute in x;
+ vm_compute in y;
+ unfold x;
+ unfold y;
+ compute [Internal.eval Internal.fold_map Internal.copy Prect]; simpl
+ >>
+ );
+ Tactics.keep ids
+ ] goal
+
+let aac_reflexivity = fun goal ->
+ aac_conclude
+ (fun zero lift ir t t' ->
+ let x,r = AAC_coq.Relation.split (lift.r) in
+ let r_reflexive = (AAC_coq.Classes.mk_reflexive x r) in
+ AAC_coq.cps_resolve_one_typeclass ~error:"The goal's relation is not reflexive"
+ r_reflexive
+ (fun reflexive ->
+ let lift_reflexivity =
+ mkApp (Lazy.force (AAC_theory.Stubs.lift_reflexivity),
+ [|
+ x;
+ r;
+ lift.e.AAC_coq.Equivalence.eq;
+ lift.lift;
+ reflexive
+ |])
+ in
+ Tacticals.tclTHEN
+ (Tactics.apply lift_reflexivity)
+ (fun goal ->
+ let concl = Tacmach.pf_concl goal in
+ let _ = pr_constr "concl "concl in
+ by_aac_reflexivity zero lift.e ir t t' goal)
+ )
+ ) goal
+
+(** A sub-tactic to lift the rewriting using AAC_lift *)
+let lift_transitivity in_left (step:constr) preorder lifting (using_eq : AAC_coq.Equivalence.t): tactic =
+ fun goal ->
+ (* catch the equation and the two members*)
+ let concl = Tacmach.pf_concl goal in
+ let (left, right, _ ) = match AAC_coq.match_as_equation goal concl with
+ | Some x -> x
+ | None -> AAC_coq.user_error "The goal is not an equation"
+ in
+ let lift_transitivity =
+ let thm =
+ if in_left
+ then
+ Lazy.force AAC_theory.Stubs.lift_transitivity_left
+ else
+ Lazy.force AAC_theory.Stubs.lift_transitivity_right
+ in
+ mkApp (thm,
+ [|
+ preorder.AAC_coq.Relation.carrier;
+ preorder.AAC_coq.Relation.r;
+ using_eq.AAC_coq.Equivalence.eq;
+ lifting;
+ step;
+ left;
+ right;
+ |])
+ in
+ Tacticals.tclTHENLIST
+ [
+ Tactics.apply lift_transitivity
+ ] goal
+
+
+(** The core tactic for aac_rewrite *)
+let core_aac_rewrite ?abort
+ rewinfo
+ subst
+ (by_aac_reflexivity : AAC_matcher.Terms.t -> AAC_matcher.Terms.t -> Proof_type.tactic)
+ (tr : constr) t left : tactic =
+ pr_constr "transitivity through" tr;
+ let tran_tac =
+ lift_transitivity rewinfo.in_left tr rewinfo.rlt rewinfo.lifting.lift rewinfo.eqt
+ in
+ AAC_coq.Rewrite.rewrite ?abort rewinfo.hypinfo subst (fun rew ->
+ Tacticals.tclTHENSV
+ (tac_or_exn (tran_tac) AAC_coq.anomaly "Unable to make the transitivity step")
+ (
+ if rewinfo.in_left
+ then [| by_aac_reflexivity left t ; rew |]
+ else [| by_aac_reflexivity t left ; rew |]
+ )
+ )
+
+exception NoSolutions
+
+
+(** Choose a substitution from a
+ [(int * Terms.t * Env.env AAC_search_monad.m) AAC_search_monad.m ] *)
+(* WARNING: Beware, since the printing function can change the order of the
+ printed monad, this function has to be updated accordingly *)
+let choose_subst subterm sol m=
+ try
+ let (depth,pat,envm) = match subterm with
+ | None -> (* first solution *)
+ List.nth ( List.rev (AAC_search_monad.to_list m)) 0
+ | Some x ->
+ List.nth ( List.rev (AAC_search_monad.to_list m)) x
+ in
+ let env = match sol with
+ None ->
+ List.nth ( (AAC_search_monad.to_list envm)) 0
+ | Some x -> List.nth ( (AAC_search_monad.to_list envm)) x
+ in
+ pat, env
+ with
+ | _ -> raise NoSolutions
+
+(** rewrite the constr modulo AC from left to right in the left member
+ of the goal *)
+let aac_rewrite ?abort rew ?(l2r=true) ?(show = false) ?(in_left=true) ?strict ~occ_subterm ~occ_sol ?extra : Proof_type.tactic = fun goal ->
+ let envs = AAC_theory.Trans.empty_envs () in
+ let (concl : Term.types) = Tacmach.pf_concl goal in
+ let (_,_,rlt) as concl =
+ match AAC_coq.match_as_equation goal concl with
+ | None -> AAC_coq.user_error "The goal is not an applied relation"
+ | Some (left, right, rlt) -> left,right,rlt
+ in
+ let check_type x =
+ Tacmach.pf_conv_x goal x rlt.AAC_coq.Relation.carrier
+ in
+ AAC_coq.Rewrite.get_hypinfo rew ~l2r ?check_type:(Some check_type)
+ (fun hypinfo ->
+ dispatch in_left concl hypinfo
+ (
+ fun rewinfo ->
+ let goal =
+ match extra with
+ | Some t -> AAC_theory.Trans.add_symbol goal rewinfo.morph_rlt envs t
+ | None -> goal
+ in
+ let pattern, subject, goal =
+ AAC_theory.Trans.t_of_constr goal rewinfo.morph_rlt envs
+ (rewinfo.pattern , rewinfo.subject)
+ in
+ let goal, ir = AAC_theory.Trans.ir_of_envs goal rewinfo.morph_rlt envs in
+
+ let units = AAC_theory.Trans.ir_to_units ir in
+ let m = AAC_matcher.subterm ?strict units pattern subject in
+ (* We sort the monad in increasing size of contet *)
+ let m = AAC_search_monad.sort (fun (x,_,_) (y,_,_) -> x - y) m in
+
+ if show
+ then
+ AAC_print.print rewinfo.morph_rlt ir m (hypinfo.AAC_coq.Rewrite.context)
+
+ else
+ try
+ let pat,subst = choose_subst occ_subterm occ_sol m in
+ let tr_step = AAC_matcher.Subst.instantiate subst pat in
+ let tr_step_raw = AAC_theory.Trans.raw_constr_of_t ir rewinfo.morph_rlt [] tr_step in
+
+ let conv = (AAC_theory.Trans.raw_constr_of_t ir rewinfo.morph_rlt (hypinfo.AAC_coq.Rewrite.context)) in
+ let subst = AAC_matcher.Subst.to_list subst in
+ let subst = List.map (fun (x,y) -> x, conv y) subst in
+ let by_aac_reflexivity = (by_aac_reflexivity rewinfo.subject rewinfo.eqt ir) in
+ core_aac_rewrite ?abort rewinfo subst by_aac_reflexivity tr_step_raw tr_step subject
+
+ with
+ | NoSolutions ->
+ Tacticals.tclFAIL 0
+ (Pp.str (if occ_subterm = None && occ_sol = None
+ then "No matching occurence modulo AC found"
+ else "No such solution"))
+ )
+ ) goal
+
+
+open Rewrite
+open Tacmach
+open Tacticals
+open Tacexpr
+open Tacinterp
+open Extraargs
+open Genarg
+
+let rec add k x = function
+ | [] -> [k,x]
+ | k',_ as ky::q ->
+ if k'=k then AAC_coq.user_error ("redondant argument ("^k^")")
+ else ky::add k x q
+
+let get k l = try Some (List.assoc k l) with Not_found -> None
+
+let get_lhs l = try List.assoc "in_right" l; false with Not_found -> true
+
+let aac_rewrite ~args =
+ aac_rewrite ~occ_subterm:(get "at" args) ~occ_sol:(get "subst" args) ~in_left:(get_lhs args)
+
+
+let pr_aac_args _ _ _ l =
+ List.fold_left
+ (fun acc -> function
+ | ("in_right" as s,_) -> Pp.(++) (Pp.str s) acc
+ | (k,i) -> Pp.(++) (Pp.(++) (Pp.str k) (Pp.int i)) acc
+ ) (Pp.str "") l
+
+ARGUMENT EXTEND aac_args
+TYPED AS ((string * int) list )
+PRINTED BY pr_aac_args
+| [ "at" integer(n) aac_args(q) ] -> [ add "at" n q ]
+| [ "subst" integer(n) aac_args(q) ] -> [ add "subst" n q ]
+| [ "in_right" aac_args(q) ] -> [ add "in_right" 0 q ]
+| [ ] -> [ [] ]
+END
+
+let pr_constro _ _ _ = fun b ->
+ match b with
+ | None -> Pp.str ""
+ | Some o -> Pp.str "<constr>"
+
+ARGUMENT EXTEND constro
+TYPED AS (constr option)
+PRINTED BY pr_constro
+| [ "[" constr(c) "]" ] -> [ Some c ]
+| [ ] -> [ None ]
+END
+
+TACTIC EXTEND _aac_reflexivity_
+| [ "aac_reflexivity" ] -> [ aac_reflexivity ]
+END
+
+TACTIC EXTEND _aac_normalise_
+| [ "aac_normalise" ] -> [ aac_normalise ]
+END
+
+TACTIC EXTEND _aac_rewrite_
+| [ "aac_rewrite" orient(l2r) constr(c) aac_args(args) constro(extra)] ->
+ [ fun gl -> aac_rewrite ?extra ~args ~l2r ~strict:true c gl ]
+END
+
+TACTIC EXTEND _aac_pattern_
+| [ "aac_pattern" orient(l2r) constr(c) aac_args(args) constro(extra)] ->
+ [ fun gl -> aac_rewrite ?extra ~args ~l2r ~strict:true ~abort:true c gl ]
+END
+
+TACTIC EXTEND _aac_instances_
+| [ "aac_instances" orient(l2r) constr(c) aac_args(args) constro(extra)] ->
+ [ fun gl -> aac_rewrite ?extra ~args ~l2r ~strict:true ~show:true c gl ]
+END
+
+TACTIC EXTEND _aacu_rewrite_
+| [ "aacu_rewrite" orient(l2r) constr(c) aac_args(args) constro(extra)] ->
+ [ fun gl -> aac_rewrite ?extra ~args ~l2r ~strict:false c gl ]
+END
+
+TACTIC EXTEND _aacu_pattern_
+| [ "aacu_pattern" orient(l2r) constr(c) aac_args(args) constro(extra)] ->
+ [ fun gl -> aac_rewrite ?extra ~args ~l2r ~strict:false ~abort:true c gl ]
+END
+
+TACTIC EXTEND _aacu_instances_
+| [ "aacu_instances" orient(l2r) constr(c) aac_args(args) constro(extra)] ->
+ [ fun gl -> aac_rewrite ?extra ~args ~l2r ~strict:false ~show:true c gl ]
+END
diff --git a/AAC_rewrite.mli b/AAC_rewrite.mli
new file mode 100644
index 0000000..d195075
--- /dev/null
+++ b/AAC_rewrite.mli
@@ -0,0 +1,9 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+(** Definition of the tactics, and corresponding Coq grammar entries.*)
diff --git a/AAC_search_monad.ml b/AAC_search_monad.ml
new file mode 100644
index 0000000..b7aabbd
--- /dev/null
+++ b/AAC_search_monad.ml
@@ -0,0 +1,65 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+type 'a m = | F of 'a
+ | N of 'a m list
+
+let fold (f : 'a -> 'b -> 'b) (m : 'a m) (acc : 'b) =
+ let rec aux acc = function
+ F x -> f x acc
+ | N l ->
+ (List.fold_left (fun acc x ->
+ match x with
+ | (N []) -> acc
+ | x -> aux acc x
+ ) acc l)
+ in
+ aux acc m
+
+
+let rec (>>) : 'a m -> ('a -> 'b m) -> 'b m =
+ fun m f ->
+ match m with
+ | F x -> f x
+ | N l ->
+ N (List.fold_left (fun acc x ->
+ match x with
+ | (N []) -> acc
+ | x -> (x >> f)::acc
+ ) [] l)
+
+let (>>|) (m : 'a m) (n :'a m) : 'a m = match (m,n) with
+ | N [],_ -> n
+ | _,N [] -> m
+ | F x, N l -> N (F x::l)
+ | N l, F x -> N (F x::l)
+ | x,y -> N [x;y]
+
+let return : 'a -> 'a m = fun x -> F x
+let fail : unit -> 'a m = fun () -> N []
+
+let sprint f m =
+ fold (fun x acc -> Printf.sprintf "%s\n%s" acc (f x)) m ""
+let rec count = function
+ | F _ -> 1
+ | N l -> List.fold_left (fun acc x -> acc+count x) 0 l
+
+let opt_comb f x y = match x with None -> f y | _ -> x
+
+let rec choose = function
+ | F x -> Some x
+ | N l -> List.fold_left (fun acc x ->
+ opt_comb choose acc x
+ ) None l
+
+let is_empty = fun x -> choose x = None
+
+let to_list m = (fold (fun x acc -> x::acc) m [])
+
+let sort f m =
+ N (List.map (fun x -> F x) (List.sort f (to_list m)))
diff --git a/AAC_search_monad.mli b/AAC_search_monad.mli
new file mode 100644
index 0000000..093f9d9
--- /dev/null
+++ b/AAC_search_monad.mli
@@ -0,0 +1,41 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+(** Search monad that allows to express non-deterministic algorithms
+ in a legible maner, or programs that solve combinatorial problems.
+
+ @see <http://spivey.oriel.ox.ac.uk/mike/search-jfp.pdf> the
+ inspiration of this module
+*)
+
+(** A data type that represent a collection of ['a] *)
+type 'a m
+
+ (** {2 Monadic operations} *)
+
+(** bind and return *)
+val ( >> ) : 'a m -> ('a -> 'b m) -> 'b m
+val return : 'a -> 'a m
+
+(** non-deterministic choice *)
+val ( >>| ) : 'a m -> 'a m -> 'a m
+
+(** failure *)
+val fail : unit -> 'a m
+
+(** folding through the collection *)
+val fold : ('a -> 'b -> 'b) -> 'a m -> 'b -> 'b
+
+(** {2 Derived facilities } *)
+
+val sprint : ('a -> string) -> 'a m -> string
+val count : 'a m -> int
+val choose : 'a m -> 'a option
+val to_list : 'a m -> 'a list
+val sort : ('a -> 'a -> int) -> 'a m -> 'a m
+val is_empty: 'a m -> bool
diff --git a/AAC_theory.ml b/AAC_theory.ml
new file mode 100644
index 0000000..55b9e21
--- /dev/null
+++ b/AAC_theory.ml
@@ -0,0 +1,1097 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+(** Constr from the theory of this particular development
+
+ The inner-working of this module is highly correlated with the
+ particular structure of {b AAC.v}, therefore, it should
+ be of little interest to most readers.
+*)
+open Term
+
+module Control = struct
+ let printing = true
+ let debug = false
+ let time = false
+end
+
+module Debug = AAC_helper.Debug (Control)
+open Debug
+
+(** {1 HMap : Specialized Hashtables based on constr} *)
+
+
+module Hashed_constr =
+struct
+ type t = constr
+ let equal = (=)
+ let hash = Hashtbl.hash
+end
+
+ (* TODO module HMap = Hashtbl, du coup ? *)
+module HMap = Hashtbl.Make(Hashed_constr)
+
+let ac_theory_path = ["AAC_tactics"; "AAC"]
+
+module Stubs = struct
+ let path = ac_theory_path@["Internal"]
+
+ (** The constants from the inductive type *)
+ let _Tty = lazy (AAC_coq.init_constant path "T")
+ let vTty = lazy (AAC_coq.init_constant path "vT")
+
+ let rsum = lazy (AAC_coq.init_constant path "sum")
+ let rprd = lazy (AAC_coq.init_constant path "prd")
+ let rsym = lazy (AAC_coq.init_constant path "sym")
+ let runit = lazy (AAC_coq.init_constant path "unit")
+
+ let vnil = lazy (AAC_coq.init_constant path "vnil")
+ let vcons = lazy (AAC_coq.init_constant path "vcons")
+ let eval = lazy (AAC_coq.init_constant path "eval")
+
+
+ let decide_thm = lazy (AAC_coq.init_constant path "decide")
+ let lift_normalise_thm = lazy (AAC_coq.init_constant path "lift_normalise")
+
+ let lift =
+ lazy (AAC_coq.init_constant ac_theory_path "AAC_lift")
+ let lift_proj_equivalence=
+ lazy (AAC_coq.init_constant ac_theory_path "aac_lift_equivalence")
+ let lift_transitivity_left =
+ lazy(AAC_coq.init_constant ac_theory_path "lift_transitivity_left")
+ let lift_transitivity_right =
+ lazy(AAC_coq.init_constant ac_theory_path "lift_transitivity_right")
+ let lift_reflexivity =
+ lazy(AAC_coq.init_constant ac_theory_path "lift_reflexivity")
+end
+
+module Classes = struct
+ module Associative = struct
+ let path = ac_theory_path
+ let typ = lazy (AAC_coq.init_constant path "Associative")
+ let ty (rlt : AAC_coq.Relation.t) (value : Term.constr) =
+ mkApp (Lazy.force typ, [| rlt.AAC_coq.Relation.carrier;
+ rlt.AAC_coq.Relation.r;
+ value
+ |] )
+ let infer goal rlt value =
+ let ty = ty rlt value in
+ AAC_coq.resolve_one_typeclass goal ty
+ end
+
+ module Commutative = struct
+ let path = ac_theory_path
+ let typ = lazy (AAC_coq.init_constant path "Commutative")
+ let ty (rlt : AAC_coq.Relation.t) (value : Term.constr) =
+ mkApp (Lazy.force typ, [| rlt.AAC_coq.Relation.carrier;
+ rlt.AAC_coq.Relation.r;
+ value
+ |] )
+
+ end
+
+ module Proper = struct
+ let path = ac_theory_path @ ["Internal";"Sym"]
+ let typeof = lazy (AAC_coq.init_constant path "type_of")
+ let relof = lazy (AAC_coq.init_constant path "rel_of")
+ let mk_typeof : AAC_coq.Relation.t -> int -> constr = fun rlt n ->
+ let x = rlt.AAC_coq.Relation.carrier in
+ mkApp (Lazy.force typeof, [| x ; AAC_coq.Nat.of_int n |])
+ let mk_relof : AAC_coq.Relation.t -> int -> constr = fun rlt n ->
+ let (x,r) = AAC_coq.Relation.split rlt in
+ mkApp (Lazy.force relof, [| x;r ; AAC_coq.Nat.of_int n |])
+
+ let ty rlt op ar =
+ let typeof = mk_typeof rlt ar in
+ let relof = mk_relof rlt ar in
+ AAC_coq.Classes.mk_morphism typeof relof op
+ let infer goal rlt op ar =
+ let ty = ty rlt op ar in
+ AAC_coq.resolve_one_typeclass goal ty
+ end
+
+ module Unit = struct
+ let path = ac_theory_path
+ let typ = lazy (AAC_coq.init_constant path "Unit")
+ let ty (rlt : AAC_coq.Relation.t) (value : Term.constr) (unit : Term.constr)=
+ mkApp (Lazy.force typ, [| rlt.AAC_coq.Relation.carrier;
+ rlt.AAC_coq.Relation.r;
+ value;
+ unit
+ |] )
+ end
+
+end
+
+(* Non empty lists *)
+module NEList = struct
+ let path = ac_theory_path @ ["Internal"]
+ let typ = lazy (AAC_coq.init_constant path "list")
+ let nil = lazy (AAC_coq.init_constant path "nil")
+ let cons = lazy (AAC_coq.init_constant path "cons")
+ let cons ty h t =
+ mkApp (Lazy.force cons, [| ty; h ; t |])
+ let nil ty x =
+ (mkApp (Lazy.force nil, [| ty ; x|]))
+ let rec of_list ty = function
+ | [] -> invalid_arg "NELIST"
+ | [x] -> nil ty x
+ | t::q -> cons ty t (of_list ty q)
+
+ let type_of_list ty =
+ mkApp (Lazy.force typ, [|ty|])
+end
+
+(** a [mset] is a ('a * pos) list *)
+let mk_mset ty (l : (constr * int) list) =
+ let pos = Lazy.force AAC_coq.Pos.typ in
+ let pair (x : constr) (ar : int) =
+ AAC_coq.Pair.of_pair ty pos (x, AAC_coq.Pos.of_int ar)
+ in
+ let pair_ty = AAC_coq.lapp AAC_coq.Pair.typ [| ty ; pos|] in
+ let rec aux = function
+ | [ ] -> assert false
+ | [x,ar] -> NEList.nil pair_ty (pair x ar)
+ | (t,ar)::q -> NEList.cons pair_ty (pair t ar) (aux q)
+ in
+ aux l
+
+module Sigma = struct
+ let sigma = lazy (AAC_coq.init_constant ac_theory_path "sigma")
+ let sigma_empty = lazy (AAC_coq.init_constant ac_theory_path "sigma_empty")
+ let sigma_add = lazy (AAC_coq.init_constant ac_theory_path "sigma_add")
+ let sigma_get = lazy (AAC_coq.init_constant ac_theory_path "sigma_get")
+
+ let add ty n x map =
+ mkApp (Lazy.force sigma_add,[|ty; n; x ; map|])
+ let empty ty =
+ mkApp (Lazy.force sigma_empty,[|ty |])
+ let typ ty =
+ mkApp (Lazy.force sigma, [|ty|])
+
+ let to_fun ty null map =
+ mkApp (Lazy.force sigma_get, [|ty;null;map|])
+
+ let of_list ty null l =
+ match l with
+ | (_,t)::q ->
+ let map =
+ List.fold_left
+ (fun acc (i,t) ->
+ assert (i > 0);
+ add ty (AAC_coq.Pos.of_int i) ( t) acc)
+ (empty ty)
+ q
+ in to_fun ty (t) map
+ | [] -> debug "of_list empty" ; to_fun ty (null) (empty ty)
+
+
+end
+
+
+module Sym = struct
+ type pack = {ar: Term.constr; value: Term.constr ; morph: Term.constr}
+ let path = ac_theory_path @ ["Internal";"Sym"]
+ let typ = lazy (AAC_coq.init_constant path "pack")
+ let mkPack = lazy (AAC_coq.init_constant path "mkPack")
+ let value = lazy (AAC_coq.init_constant path "value")
+ let null = lazy (AAC_coq.init_constant path "null")
+ let mk_pack (rlt: AAC_coq.Relation.t) s =
+ let (x,r) = AAC_coq.Relation.split rlt in
+ mkApp (Lazy.force mkPack, [|x;r; s.ar;s.value;s.morph|])
+ let null rlt =
+ let x = rlt.AAC_coq.Relation.carrier in
+ let r = rlt.AAC_coq.Relation.r in
+ mkApp (Lazy.force null, [| x;r;|])
+
+ let mk_ty : AAC_coq.Relation.t -> constr = fun rlt ->
+ let (x,r) = AAC_coq.Relation.split rlt in
+ mkApp (Lazy.force typ, [| x; r|] )
+end
+
+module Bin =struct
+ type pack = {value : Term.constr;
+ compat : Term.constr;
+ assoc : Term.constr;
+ comm : Term.constr option;
+ }
+
+ let path = ac_theory_path @ ["Internal";"Bin"]
+ let typ = lazy (AAC_coq.init_constant path "pack")
+ let mkPack = lazy (AAC_coq.init_constant path "mk_pack")
+
+ let mk_pack: AAC_coq.Relation.t -> pack -> Term.constr = fun (rlt) s ->
+ let (x,r) = AAC_coq.Relation.split rlt in
+ let comm_ty = Classes.Commutative.ty rlt s.value in
+ mkApp (Lazy.force mkPack , [| x ; r;
+ s.value;
+ s.compat ;
+ s.assoc;
+ AAC_coq.Option.of_option comm_ty s.comm
+ |])
+ let mk_ty : AAC_coq.Relation.t -> constr = fun rlt ->
+ let (x,r) = AAC_coq.Relation.split rlt in
+ mkApp (Lazy.force typ, [| x; r|] )
+end
+
+module Unit = struct
+ let path = ac_theory_path @ ["Internal"]
+ let unit_of_ty = lazy (AAC_coq.init_constant path "unit_of")
+ let unit_pack_ty = lazy (AAC_coq.init_constant path "unit_pack")
+ let mk_unit_of = lazy (AAC_coq.init_constant path "mk_unit_for")
+ let mk_unit_pack = lazy (AAC_coq.init_constant path "mk_unit_pack")
+
+ type unit_of =
+ {
+ uf_u : Term.constr; (* u *)
+ uf_idx : Term.constr;
+ uf_desc : Term.constr; (* Unit R (e_bin uf_idx) u *)
+ }
+
+ type pack = {
+ u_value : Term.constr; (* X *)
+ u_desc : (unit_of) list (* unit_of u_value *)
+ }
+
+ let ty_unit_of rlt e_bin u =
+ let (x,r) = AAC_coq.Relation.split rlt in
+ mkApp ( Lazy.force unit_of_ty, [| x; r; e_bin; u |])
+
+ let ty_unit_pack rlt e_bin =
+ let (x,r) = AAC_coq.Relation.split rlt in
+ mkApp (Lazy.force unit_pack_ty, [| x; r; e_bin |])
+
+ let mk_unit_of rlt e_bin u unit_of =
+ let (x,r) = AAC_coq.Relation.split rlt in
+ mkApp (Lazy.force mk_unit_of , [| x;
+ r;
+ e_bin ;
+ u;
+ unit_of.uf_idx;
+ unit_of.uf_desc
+ |])
+
+ let mk_pack rlt e_bin pack : Term.constr =
+ let (x,r) = AAC_coq.Relation.split rlt in
+ let ty = ty_unit_of rlt e_bin pack.u_value in
+ let mk_unit_of = mk_unit_of rlt e_bin pack.u_value in
+ let u_desc =AAC_coq.List.of_list ( ty ) (List.map mk_unit_of pack.u_desc) in
+ mkApp (Lazy.force mk_unit_pack, [|x;r;
+ e_bin ;
+ pack.u_value;
+ u_desc
+ |])
+
+ let default x : pack =
+ { u_value = x;
+ u_desc = []
+ }
+
+end
+
+let anomaly msg =
+ Util.anomaly ("aac_tactics: " ^ msg)
+
+let user_error msg =
+ Util.error ("aac_tactics: " ^ msg)
+
+module Trans = struct
+
+ (** {1 From Coq to Abstract Syntax Trees (AST)}
+
+ This module provides facilities to interpret a Coq term with
+ arbitrary operators as an abstract syntax tree. Considering an
+ application, we try to infer instances of our classes.
+
+ We consider that [A] operators are coarser than [AC] operators,
+ which in turn are coarser than raw morphisms. That means that
+ [List.append], of type [(A : Type) -> list A -> list A -> list
+ A] can be understood as an [A] operator.
+
+ During this reification, we gather some informations that will
+ be used to rebuild Coq terms from AST ( type {!envs})
+
+ We use a main hash-table from [constr] to [pack], in order to
+ discriminate the various constructors. All these are mixed in
+ order to improve on the number of comparisons in the tables *)
+
+ (* 'a * (unit, op_unit) option *)
+ type 'a with_unit = 'a * (Unit.unit_of) option
+ type pack =
+ (* used to infer the AC/A structure in the first pass {!Gather} *)
+ | Bin of Bin.pack with_unit
+ (* will only be used in the second pass : {!Parse}*)
+ | Sym of Sym.pack
+ | Unit of constr (* to avoid confusion in bloom *)
+
+ (** discr is a map from {!Term.constr} to {!pack}.
+
+ [None] means that it is not possible to instantiate this partial
+ application to an interesting class.
+
+ [Some x] means that we found something in the past. This means
+ in particular that a single [constr] cannot be two things at the
+ same time.
+
+ The field [bloom] allows to give a unique number to each class we
+ found. *)
+ type envs =
+ {
+ discr : (pack option) HMap.t ;
+ bloom : (pack, int ) Hashtbl.t;
+ bloom_back : (int, pack ) Hashtbl.t;
+ bloom_next : int ref;
+ }
+
+ let empty_envs () =
+ {
+ discr = HMap.create 17;
+ bloom = Hashtbl.create 17;
+ bloom_back = Hashtbl.create 17;
+ bloom_next = ref 1;
+ }
+
+
+
+ let add_bloom envs pack =
+ if Hashtbl.mem envs.bloom pack
+ then ()
+ else
+ let x = ! (envs.bloom_next) in
+ Hashtbl.add envs.bloom pack x;
+ Hashtbl.add envs.bloom_back x pack;
+ incr (envs.bloom_next)
+
+ let find_bloom envs pack =
+ try Hashtbl.find envs.bloom pack
+ with Not_found -> assert false
+
+ (********************************************)
+ (* (\* Gather the occuring AC/A symbols *\) *)
+ (********************************************)
+
+ (** This modules exhibit a function that memoize in the environment
+ all the AC/A operators as well as the morphism that occur. This
+ staging process allows us to prefer AC/A operators over raw
+ morphisms. Moreover, for each AC/A operators, we need to try to
+ infer units. Otherwise, we do not have [x * y * x <= a * a] since 1
+ does not occur.
+
+ But, do we also need to check whether constants are
+ units. Otherwise, we do not have the ability to rewrite [0 = a +
+ a] in [a = ...]*)
+ module Gather : sig
+ val gather : AAC_coq.goal_sigma -> AAC_coq.Relation.t -> envs -> Term.constr -> AAC_coq.goal_sigma
+ end
+ = struct
+
+ let memoize envs t pack : unit =
+ begin
+ HMap.add envs.discr t (Some pack);
+ add_bloom envs pack;
+ match pack with
+ | Bin (_, None) | Sym _ -> ()
+ | Bin (_, Some (unit_of)) ->
+ let unit = unit_of.Unit.uf_u in
+ HMap.add envs.discr unit (Some (Unit unit));
+ add_bloom envs (Unit unit);
+ | Unit _ -> assert false
+ end
+
+
+ let get_unit (rlt : AAC_coq.Relation.t) op goal :
+ (AAC_coq.goal_sigma * constr * constr ) option=
+ let x = (rlt.AAC_coq.Relation.carrier) in
+ let unit, goal = AAC_coq.evar_unit goal x in
+ let ty =Classes.Unit.ty rlt op unit in
+ let result =
+ try
+ let t,goal = AAC_coq.resolve_one_typeclass goal ty in
+ Some (goal,t,unit)
+ with Not_found -> None
+ in
+ match result with
+ | None -> None
+ | Some (goal,s,unit) ->
+ let unit = AAC_coq.nf_evar goal unit in
+ Some (goal, unit, s)
+
+
+
+ (** gives back the class and the operator *)
+ let is_bin (rlt: AAC_coq.Relation.t) (op: constr) ( goal: AAC_coq.goal_sigma)
+ : (AAC_coq.goal_sigma * Bin.pack) option =
+ let assoc_ty = Classes.Associative.ty rlt op in
+ let comm_ty = Classes.Commutative.ty rlt op in
+ let proper_ty = Classes.Proper.ty rlt op 2 in
+ try
+ let proper , goal = AAC_coq.resolve_one_typeclass goal proper_ty in
+ let assoc, goal = AAC_coq.resolve_one_typeclass goal assoc_ty in
+ let comm , goal =
+ try
+ let comm, goal = AAC_coq.resolve_one_typeclass goal comm_ty in
+ Some comm, goal
+ with Not_found ->
+ None, goal
+ in
+ let bin =
+ {Bin.value = op;
+ Bin.compat = proper;
+ Bin.assoc = assoc;
+ Bin.comm = comm }
+ in
+ Some (goal,bin)
+ with |Not_found -> None
+
+ let is_bin (rlt : AAC_coq.Relation.t) (op : constr) (goal : AAC_coq.goal_sigma)=
+ match is_bin rlt op goal with
+ | None -> None
+ | Some (goal, bin_pack) ->
+ match get_unit rlt op goal with
+ | None -> Some (goal, Bin (bin_pack, None))
+ | Some (gl, unit,s) ->
+ let unit_of =
+ {
+ Unit.uf_u = unit;
+ (* TRICK : this term is not well-typed by itself,
+ but it is okay the way we use it in the other
+ functions *)
+ Unit.uf_idx = op;
+ Unit.uf_desc = s;
+ }
+ in Some (gl,Bin (bin_pack, Some (unit_of)))
+
+
+ (** {is_morphism} try to infer the kind of operator we are
+ dealing with *)
+ let is_morphism goal (rlt : AAC_coq.Relation.t) (papp : Term.constr) (ar : int) : (AAC_coq.goal_sigma * pack ) option =
+ let typeof = Classes.Proper.mk_typeof rlt ar in
+ let relof = Classes.Proper.mk_relof rlt ar in
+ let m = AAC_coq.Classes.mk_morphism typeof relof papp in
+ try
+ let m,goal = AAC_coq.resolve_one_typeclass goal m in
+ let pack = {Sym.ar = (AAC_coq.Nat.of_int ar); Sym.value= papp; Sym.morph= m} in
+ Some (goal, Sym pack)
+ with
+ | Not_found -> None
+
+
+ (* [crop_app f [| a_0 ; ... ; a_n |]]
+ returns Some (f a_0 ... a_(n-2), [|a_(n-1); a_n |] )
+ *)
+ let crop_app t ca : (constr * constr array) option=
+ let n = Array.length ca in
+ if n <= 1
+ then None
+ else
+ let papp = Term.mkApp (t, Array.sub ca 0 (n-2) ) in
+ let args = Array.sub ca (n-2) 2 in
+ Some (papp, args )
+
+ let fold goal (rlt : AAC_coq.Relation.t) envs t ca cont =
+ let fold_morphism t ca =
+ let nb_params = Array.length ca in
+ let rec aux n =
+ assert (n < nb_params && 0 < nb_params );
+ let papp = Term.mkApp (t, Array.sub ca 0 n) in
+ match is_morphism goal rlt papp (nb_params - n) with
+ | None ->
+ (* here we have to memoize the failures *)
+ HMap.add envs.discr papp None;
+ if n < nb_params - 1 then aux (n+1) else goal
+ | Some (goal, pack) ->
+ let args = Array.sub ca (n) (nb_params -n)in
+ let goal = Array.fold_left cont goal args in
+ memoize envs papp pack;
+ goal
+ in
+ if nb_params = 0 then goal else aux 0
+ in
+ let is_aac t = is_bin rlt t in
+ match crop_app t ca with
+ | None ->
+ fold_morphism t ca
+ | Some (papp, args) ->
+ begin match is_aac papp goal with
+ | None -> fold_morphism t ca
+ | Some (goal, pack) ->
+ memoize envs papp pack;
+ Array.fold_left cont goal args
+ end
+
+ (* update in place the envs of known stuff, using memoization. We
+ have to memoize failures, here. *)
+ let gather goal (rlt : AAC_coq.Relation.t ) envs t : AAC_coq.goal_sigma =
+ let rec aux goal x =
+ match AAC_coq.decomp_term x with
+ | App (t,ca) ->
+ fold goal rlt envs t ca (aux )
+ | _ -> goal
+ in
+ aux goal t
+ end
+
+ (** We build a term out of a constr, updating in place the
+ environment if needed (the only kind of such updates are the
+ constants). *)
+ module Parse :
+ sig
+ val t_of_constr : AAC_coq.goal_sigma -> AAC_coq.Relation.t -> envs -> constr -> AAC_matcher.Terms.t * AAC_coq.goal_sigma
+ end
+ = struct
+
+ (** [discriminates goal envs rlt t ca] infer :
+
+ - its {! pack } (is it an AC operator, or an A operator, or a
+ Symbol ?),
+
+ - the relevant partial application,
+
+ - the vector of the remaining arguments
+
+ We use an expansion to handle the special case of units,
+ before going back to the general discrimination
+ procedure. That means that a unit is more coarse than a raw
+ morphism
+
+ This functions is prevented to go through [ar < 0] by the fact
+ that a constant is a morphism. But not an eva. *)
+
+ let is_morphism goal (rlt : AAC_coq.Relation.t) (papp : Term.constr) (ar : int) : (AAC_coq.goal_sigma * pack ) option =
+ let typeof = Classes.Proper.mk_typeof rlt ar in
+ let relof = Classes.Proper.mk_relof rlt ar in
+ let m = AAC_coq.Classes.mk_morphism typeof relof papp in
+ try
+ let m,goal = AAC_coq.resolve_one_typeclass goal m in
+ let pack = {Sym.ar = (AAC_coq.Nat.of_int ar); Sym.value= papp; Sym.morph= m} in
+ Some (goal, Sym pack)
+ with
+ | e -> None
+
+ exception NotReflexive
+ let discriminate goal envs (rlt : AAC_coq.Relation.t) t ca : AAC_coq.goal_sigma * pack * constr * constr array =
+ let nb_params = Array.length ca in
+ let rec fold goal ar :AAC_coq.goal_sigma * pack * constr * constr array =
+ begin
+ assert (0 <= ar && ar <= nb_params);
+ let p_app = mkApp (t, Array.sub ca 0 (nb_params - ar)) in
+ begin
+ try
+ begin match HMap.find envs.discr p_app with
+ | None ->
+ fold goal (ar-1)
+ | Some pack ->
+ (goal, pack, p_app, Array.sub ca (nb_params -ar ) ar)
+ end
+ with
+ Not_found -> (* Could not find this constr *)
+ memoize (is_morphism goal rlt p_app ar) p_app ar
+ end
+ end
+ and memoize (x) p_app ar =
+ assert (0 <= ar && ar <= nb_params);
+ match x with
+ | Some (goal, pack) ->
+ HMap.add envs.discr p_app (Some pack);
+ add_bloom envs pack;
+ (goal, pack, p_app, Array.sub ca (nb_params-ar) ar)
+ | None ->
+
+ if ar = 0 then raise NotReflexive;
+ begin
+ (* to memoise the failures *)
+ HMap.add envs.discr p_app None;
+ (* will terminate, since [const] is capped, and it is
+ easy to find an instance of a constant *)
+ fold goal (ar-1)
+ end
+ in
+ try match HMap.find envs.discr (mkApp (t,ca))with
+ | None -> fold goal (nb_params)
+ | Some pack -> goal, pack, (mkApp (t,ca)), [| |]
+ with Not_found -> fold goal (nb_params)
+
+ let discriminate goal envs rlt x =
+ try
+ match AAC_coq.decomp_term x with
+ | App (t,ca) ->
+ discriminate goal envs rlt t ca
+ | _ -> discriminate goal envs rlt x [| |]
+ with
+ | NotReflexive -> user_error "The relation to which the goal was lifted is not Reflexive"
+ (* TODO: is it the only source of invalid use that fall
+ into this catch_all ? *)
+ | e ->
+ user_error "Cannot handle this kind of hypotheses (variables occuring under a function symbol which is not a proper morphism)."
+
+ (** [t_of_constr goal rlt envs cstr] builds the abstract syntax tree
+ of the term [cstr]. Doing so, it modifies the environment of
+ known stuff [envs], and eventually creates fresh
+ evars. Therefore, we give back the goal updated accordingly *)
+ let t_of_constr goal (rlt: AAC_coq.Relation.t ) envs : constr -> AAC_matcher.Terms.t * AAC_coq.goal_sigma =
+ let r_goal = ref (goal) in
+ let rec aux x =
+ match AAC_coq.decomp_term x with
+ | Rel i -> AAC_matcher.Terms.Var i
+ | _ ->
+ let goal, pack , p_app, ca = discriminate (!r_goal) envs rlt x in
+ r_goal := goal;
+ let k = find_bloom envs pack
+ in
+ match pack with
+ | Bin (pack,_) ->
+ begin match pack.Bin.comm with
+ | Some _ ->
+ assert (Array.length ca = 2);
+ AAC_matcher.Terms.Plus ( k, aux ca.(0), aux ca.(1))
+ | None ->
+ assert (Array.length ca = 2);
+ AAC_matcher.Terms.Dot ( k, aux ca.(0), aux ca.(1))
+ end
+ | Unit _ ->
+ assert (Array.length ca = 0);
+ AAC_matcher.Terms.Unit ( k)
+ | Sym _ ->
+ AAC_matcher.Terms.Sym ( k, Array.map aux ca)
+ in
+ (
+ fun x -> let r = aux x in r, !r_goal
+ )
+
+ end (* Parse *)
+
+ let add_symbol goal rlt envs l =
+ let goal = Gather.gather goal rlt envs (Term.mkApp (l, [| Term.mkRel 0;Term.mkRel 0|])) in
+ goal
+
+ (* [t_of_constr] buils a the abstract syntax tree of a constr,
+ updating in place the environment. Doing so, we infer all the
+ morphisms and the AC/A operators. It is mandatory to do so both
+ for the pattern and the term, since AC symbols can occur in one
+ and not the other *)
+ let t_of_constr goal rlt envs (l,r) =
+ let goal = Gather.gather goal rlt envs l in
+ let goal = Gather.gather goal rlt envs r in
+ let l,goal = Parse.t_of_constr goal rlt envs l in
+ let r, goal = Parse.t_of_constr goal rlt envs r in
+ l, r, goal
+
+ (* An intermediate representation of the environment, with association lists for AC/A operators, and also the raw [packer] information *)
+
+ type ir =
+ {
+ packer : (int, pack) Hashtbl.t; (* = bloom_back *)
+ bin : (int * Bin.pack) list ;
+ units : (int * Unit.pack) list;
+ sym : (int * Term.constr) list ;
+ matcher_units : AAC_matcher.ext_units
+ }
+
+ let ir_to_units ir = ir.matcher_units
+
+ let ir_of_envs goal (rlt : AAC_coq.Relation.t) envs =
+ let add x y l = (x,y)::l in
+ let nil = [] in
+ let sym ,
+ (bin : (int * Bin.pack with_unit) list),
+ (units : (int * constr) list) =
+ Hashtbl.fold
+ (fun int pack (sym,bin,units) ->
+ match pack with
+ | Bin s ->
+ sym, add (int) s bin, units
+ | Sym s ->
+ add (int) s sym, bin, units
+ | Unit s ->
+ sym, bin, add int s units
+ )
+ envs.bloom_back
+ (nil,nil,nil)
+ in
+ let matcher_units =
+ let unit_for , is_ac =
+ List.fold_left
+ (fun ((unit_for,is_ac) as acc) (n,(bp,wu)) ->
+ match wu with
+ | None -> acc
+ | Some (unit_of) ->
+ let unit_n = Hashtbl.find envs.bloom
+ (Unit unit_of.Unit.uf_u)
+ in
+ ( n, unit_n)::unit_for,
+ (n, bp.Bin.comm <> None )::is_ac
+
+ )
+ ([],[]) bin
+ in
+ {AAC_matcher.unit_for = unit_for; AAC_matcher.is_ac = is_ac}
+
+ in
+ let units : (int * Unit.pack) list =
+ List.fold_left
+ (fun acc (n,u) ->
+ (* first, gather all bins with this unit *)
+ let all_bin =
+ List.fold_left
+ ( fun acc (nop,(bp,wu)) ->
+ match wu with
+ | None -> acc
+ | Some unit_of ->
+ if unit_of.Unit.uf_u = u
+ then
+ {unit_of with
+ Unit.uf_idx = (AAC_coq.Pos.of_int nop)} :: acc
+ else
+ acc
+ )
+ [] bin
+ in
+ (n,{
+ Unit.u_value = u;
+ Unit.u_desc = all_bin
+ })::acc
+ )
+ [] units
+
+ in
+ goal, {
+ packer = envs.bloom_back;
+ bin = (List.map (fun (n,(s,_)) -> n, s) bin);
+ units = units;
+ sym = (List.map (fun (n,s) -> n,(Sym.mk_pack rlt s)) sym);
+ matcher_units = matcher_units
+ }
+
+
+
+ (** {1 From AST back to Coq }
+
+ The next functions allow one to map OCaml abstract syntax trees
+ to Coq terms *)
+
+ (** {2 Building raw, natural, terms} *)
+
+ (** [raw_constr_of_t_debruijn] rebuilds a term in the raw
+ representation, without products on top, and maybe escaping free
+ debruijn indices (in the case of a pattern for example). *)
+ let raw_constr_of_t_debruijn ir (t : AAC_matcher.Terms.t) : Term.constr * int list =
+ let add_set,get =
+ let r = ref [] in
+ let rec add x = function
+ [ ] -> [x]
+ | t::q when t = x -> t::q
+ | t::q -> t:: (add x q)
+ in
+ (fun x -> r := add x !r),(fun () -> !r)
+ in
+ (* Here, we rely on the invariant that the maps are well formed:
+ it is meanigless to fail to find a symbol in the maps, or to
+ find the wrong kind of pack in the maps *)
+ let rec aux t =
+ match t with
+ | AAC_matcher.Terms.Plus (s,l,r) ->
+ begin match Hashtbl.find ir.packer s with
+ | Bin (s,_) ->
+ mkApp (s.Bin.value , [|(aux l);(aux r)|])
+ | _ -> Printf.printf "erreur:%i\n%!"s;
+ assert false
+ end
+ | AAC_matcher.Terms.Dot (s,l,r) ->
+ begin match Hashtbl.find ir.packer s with
+ | Bin (s,_) ->
+ mkApp (s.Bin.value, [|(aux l);(aux r)|])
+ | _ -> assert false
+ end
+ | AAC_matcher.Terms.Sym (s,t) ->
+ begin match Hashtbl.find ir.packer s with
+ | Sym s ->
+ mkApp (s.Sym.value, Array.map aux t)
+ | _ -> assert false
+ end
+ | AAC_matcher.Terms.Unit x ->
+ begin match Hashtbl.find ir.packer x with
+ | Unit s -> s
+ | _ -> assert false
+ end
+ | AAC_matcher.Terms.Var i -> add_set i;
+ mkRel (i)
+ in
+ let t = aux t in
+ t , get ( )
+
+ (** [raw_constr_of_t] rebuilds a term in the raw representation *)
+ let raw_constr_of_t ir rlt (context:rel_context) t =
+ (** cap rebuilds the products in front of the constr *)
+ let rec cap c = function [] -> c
+ | t::q ->
+ let i = Term.lookup_rel t context in
+ cap (mkProd_or_LetIn i c) q
+ in
+ let t,indices = raw_constr_of_t_debruijn ir t in
+ cap t (List.sort (Pervasives.compare) indices)
+
+
+ (** {2 Building reified terms} *)
+
+ (* Some informations to be able to build the maps *)
+ type reif_params =
+ {
+ bin_null : Bin.pack; (* the default A operator *)
+ sym_null : constr;
+ unit_null : Unit.pack;
+ sym_ty : constr; (* the type, as it appears in e_sym *)
+ bin_ty : constr
+ }
+
+
+ (** A record containing the environments that will be needed by the
+ decision procedure, as a Coq constr. Contains the functions
+ from the symbols (as ints) to indexes (as constr) *)
+
+ type sigmas = {
+ env_sym : Term.constr;
+ env_bin : Term.constr;
+ env_units : Term.constr; (* the [idx -> X:constr] *)
+ }
+
+
+ type sigma_maps = {
+ sym_to_pos : int -> Term.constr;
+ bin_to_pos : int -> Term.constr;
+ units_to_pos : int -> Term.constr;
+ }
+
+
+ (** infers some stuff that will be required when we will build
+ environments (our environments need a default case, so we need
+ an Op_AC, an Op_A, and a symbol) *)
+ (* Note : this function can fail if the user is using the wrong
+ relation, like proving a = b, while the classes are defined with
+ another relation (==) *)
+ let build_reif_params goal (rlt : AAC_coq.Relation.t) (zero) :
+ reif_params * AAC_coq.goal_sigma =
+ let carrier = rlt.AAC_coq.Relation.carrier in
+ let bin_null =
+ try
+ let op,goal = AAC_coq.evar_binary goal carrier in
+ let assoc,goal = Classes.Associative.infer goal rlt op in
+ let compat,goal = Classes.Proper.infer goal rlt op 2 in
+ let op = AAC_coq.nf_evar goal op in
+ {
+ Bin.value = op;
+ Bin.compat = compat;
+ Bin.assoc = assoc;
+ Bin.comm = None
+ }
+ with Not_found -> user_error "Cannot infer a default A operator (required at least to be Proper and Associative)"
+ in
+ let zero, goal =
+ try
+ let evar_op,goal = AAC_coq.evar_binary goal carrier in
+ let evar_unit, goal = AAC_coq.evar_unit goal carrier in
+ let query = Classes.Unit.ty rlt evar_op evar_unit in
+ let _, goal = AAC_coq.resolve_one_typeclass goal query in
+ AAC_coq.nf_evar goal evar_unit, goal
+ with _ -> zero, goal in
+ let sym_null = Sym.null rlt in
+ let unit_null = Unit.default zero in
+ let record =
+ {
+ bin_null = bin_null;
+ sym_null = sym_null;
+ unit_null = unit_null;
+ sym_ty = Sym.mk_ty rlt ;
+ bin_ty = Bin.mk_ty rlt
+ }
+ in
+ record, goal
+
+ (* We want to lift down the indexes of symbols. *)
+ let renumber (l: (int * 'a) list ) =
+ let _, l = List.fold_left (fun (next,acc) (glob,x) ->
+ (next+1, (next,glob,x)::acc)
+ ) (1,[]) l
+ in
+ let rec to_global loc = function
+ | [] -> assert false
+ | (local, global,x)::q when local = loc -> global
+ | _::q -> to_global loc q
+ in
+ let rec to_local glob = function
+ | [] -> assert false
+ | (local, global,x)::q when global = glob -> local
+ | _::q -> to_local glob q
+ in
+ let locals = List.map (fun (local,global,x) -> local,x) l in
+ locals, (fun x -> to_local x l) , (fun x -> to_global x l)
+
+ (** [build_sigma_maps] given a envs and some reif_params, we are
+ able to build the sigmas *)
+ let build_sigma_maps (rlt : AAC_coq.Relation.t) zero ir (k : sigmas * sigma_maps -> Proof_type.tactic ) : Proof_type.tactic = fun goal ->
+ let rp,goal = build_reif_params goal rlt zero in
+ let renumbered_sym, to_local, to_global = renumber ir.sym in
+ let env_sym = Sigma.of_list
+ rp.sym_ty
+ (rp.sym_null)
+ renumbered_sym
+ in
+ AAC_coq.cps_mk_letin "env_sym" env_sym
+ (fun env_sym ->
+ let bin = (List.map ( fun (n,s) -> n, Bin.mk_pack rlt s) ir.bin) in
+ let env_bin =
+ Sigma.of_list
+ rp.bin_ty
+ (Bin.mk_pack rlt rp.bin_null)
+ bin
+ in
+ AAC_coq.cps_mk_letin "env_bin" env_bin
+ (fun env_bin ->
+ let units = (List.map (fun (n,s) -> n, Unit.mk_pack rlt env_bin s)ir.units) in
+ let env_units =
+ Sigma.of_list
+ (Unit.ty_unit_pack rlt env_bin)
+ (Unit.mk_pack rlt env_bin rp.unit_null )
+ units
+ in
+
+ AAC_coq.cps_mk_letin "env_units" env_units
+ (fun env_units ->
+ let sigmas =
+ {
+ env_sym = env_sym ;
+ env_bin = env_bin ;
+ env_units = env_units;
+ } in
+ let f = List.map (fun (x,_) -> (x,AAC_coq.Pos.of_int x)) in
+ let sigma_maps =
+ {
+ sym_to_pos = (let sym = f renumbered_sym in fun x -> (List.assoc (to_local x) sym));
+ bin_to_pos = (let bin = f bin in fun x -> (List.assoc x bin));
+ units_to_pos = (let units = f units in fun x -> (List.assoc x units));
+ }
+ in
+ k (sigmas, sigma_maps )
+ )
+ )
+ ) goal
+
+ (** builders for the reification *)
+ type reif_builders =
+ {
+ rsum: constr -> constr -> constr -> constr;
+ rprd: constr -> constr -> constr -> constr;
+ rsym: constr -> constr array -> constr;
+ runit : constr -> constr
+ }
+
+ (* donne moi une tactique, je rajoute ma part. Potentiellement, il
+ est possible d'utiliser la notation 'do' a la Haskell:
+ http://www.cas.mcmaster.ca/~carette/pa_monad/ *)
+ let (>>) : 'a * Proof_type.tactic -> ('a -> 'b * Proof_type.tactic) -> 'b * Proof_type.tactic =
+ fun (x,t) f ->
+ let b,t' = f x in
+ b, Tacticals.tclTHEN t t'
+
+ let return x = x, Tacticals.tclIDTAC
+
+ let mk_vect vnil vcons v =
+ let ar = Array.length v in
+ let rec aux = function
+ | 0 -> vnil
+ | n -> let n = n-1 in
+ mkApp( vcons,
+ [|
+ (AAC_coq.Nat.of_int n);
+ v.(ar - 1 - n);
+ (aux (n))
+ |]
+ )
+ in aux ar
+
+ (* TODO: use a do notation *)
+ let mk_reif_builders (rlt: AAC_coq.Relation.t) (env_sym:constr) (k: reif_builders -> Proof_type.tactic) =
+ let x = (rlt.AAC_coq.Relation.carrier) in
+ let r = (rlt.AAC_coq.Relation.r) in
+
+ let x_r_env = [|x;r;env_sym|] in
+ let tty = mkApp (Lazy.force Stubs._Tty, x_r_env) in
+ let rsum = mkApp (Lazy.force Stubs.rsum, x_r_env) in
+ let rprd = mkApp (Lazy.force Stubs.rprd, x_r_env) in
+ let rsym = mkApp (Lazy.force Stubs.rsym, x_r_env) in
+ let vnil = mkApp (Lazy.force Stubs.vnil, x_r_env) in
+ let vcons = mkApp (Lazy.force Stubs.vcons, x_r_env) in
+ AAC_coq.cps_mk_letin "tty" tty
+ (fun tty ->
+ AAC_coq.cps_mk_letin "rsum" rsum
+ (fun rsum ->
+ AAC_coq.cps_mk_letin "rprd" rprd
+ (fun rprd ->
+ AAC_coq.cps_mk_letin "rsym" rsym
+ (fun rsym ->
+ AAC_coq.cps_mk_letin "vnil" vnil
+ (fun vnil ->
+ AAC_coq.cps_mk_letin "vcons" vcons
+ (fun vcons ->
+ let r ={
+ rsum =
+ begin fun idx l r ->
+ mkApp (rsum, [| idx ; mk_mset tty [l,1 ; r,1]|])
+ end;
+ rprd =
+ begin fun idx l r ->
+ let lst = NEList.of_list tty [l;r] in
+ mkApp (rprd, [| idx; lst|])
+ end;
+ rsym =
+ begin fun idx v ->
+ let vect = mk_vect vnil vcons v in
+ mkApp (rsym, [| idx; vect|])
+ end;
+ runit = fun idx -> (* could benefit of a letin *)
+ mkApp (Lazy.force Stubs.runit , [|x;r;env_sym;idx; |])
+ }
+ in k r
+ ))))))
+
+
+
+ type reifier = sigma_maps * reif_builders
+
+
+ let mk_reifier rlt zero envs (k : sigmas *reifier -> Proof_type.tactic) =
+ build_sigma_maps rlt zero envs
+ (fun (s,sm) ->
+ mk_reif_builders rlt s.env_sym
+ (fun rb ->k (s,(sm,rb)) )
+
+ )
+
+ (** [reif_constr_of_t reifier t] rebuilds the term [t] in the
+ reified form. We use the [reifier] to minimise the size of the
+ terms (we make as much lets as possible)*)
+ let reif_constr_of_t (sm,rb) (t:AAC_matcher.Terms.t) : constr =
+ let rec aux = function
+ | AAC_matcher.Terms.Plus (s,l,r) ->
+ let idx = sm.bin_to_pos s in
+ rb.rsum idx (aux l) (aux r)
+ | AAC_matcher.Terms.Dot (s,l,r) ->
+ let idx = sm.bin_to_pos s in
+ rb.rprd idx (aux l) (aux r)
+ | AAC_matcher.Terms.Sym (s,t) ->
+ let idx = sm.sym_to_pos s in
+ rb.rsym idx (Array.map aux t)
+ | AAC_matcher.Terms.Unit s ->
+ let idx = sm.units_to_pos s in
+ rb.runit idx
+ | AAC_matcher.Terms.Var i ->
+ anomaly "call to reif_constr_of_t on a term with variables."
+ in aux t
+end
+
+
+
diff --git a/theory.mli b/AAC_theory.mli
index 85ad24b..2f57446 100644
--- a/theory.mli
+++ b/AAC_theory.mli
@@ -7,8 +7,8 @@
(***************************************************************************)
(** Bindings for Coq constants that are specific to the plugin;
- reification and translation functions.
-
+ reification and translation functions.
+
Note: this module is highly correlated with the definitions of {i
AAC.v}.
@@ -20,15 +20,22 @@
AAC.v}).
*)
+(** Both in OCaml and Coq, we represent finite multisets using
+ weighted lists ([('a*int) list]), see {!AAC_matcher.mset}.
+
+ [mk_mset ty l] constructs a Coq multiset from an OCaml multiset
+ [l] of Coq terms of type [ty] *)
+val mk_mset:Term.constr -> (Term.constr * int) list ->Term.constr
+(** {2 Packaging modules} *)
(** Environments *)
module Sigma:
sig
(** [add ty n x map] adds the value [x] of type [ty] with key [n] in [map] *)
val add: Term.constr ->Term.constr ->Term.constr ->Term.constr ->Term.constr
-
+
(** [empty ty] create an empty map of type [ty] *)
val empty: Term.constr ->Term.constr
@@ -39,19 +46,6 @@ sig
val to_fun: Term.constr ->Term.constr ->Term.constr ->Term.constr
end
-(** The constr associated with our typeclasses for AC or A operators *)
-val op_ac:Term.constr lazy_t
-val op_a:Term.constr lazy_t
-
-(** The evaluation fonction, used to convert a reified coq term to a
- raw coq term *)
-val eval: Term.constr lazy_t
-
-(** The main lemma of our theory, that is
- [compare (norm u) (norm v) = Eq -> eval u == eval v] *)
-val decide_thm:Term.constr lazy_t
-
-(** {2 Packaging modules} *)
(** Dynamically typed morphisms *)
module Sym:
@@ -61,80 +55,48 @@ sig
val typ: Term.constr lazy_t
- (** [mk_typeof rlt int] *)
- val mk_typeof: Coq.reltype ->int ->Term.constr
-
- (** [mk_typeof rlt int] *)
- val mk_relof: Coq.reltype ->int ->Term.constr
(** [mk_pack rlt (ar,value,morph)] *)
- val mk_pack: Coq.reltype -> pack -> Term.constr
-
- (** [null] builds a dummy symbol, given an {!Coq.eqtype} and a
- constant *)
- val null: Coq.eqtype -> Term.constr -> Term.constr
-
-end
-
-(** Dynamically typed AC operations *)
-module Sum:
-sig
-
- (** mimics the Coq record [Sum.pack] *)
- type pack = {plus : Term.constr; zero: Term.constr ; opac: Term.constr}
-
- val typ: Term.constr lazy_t
-
- (** [mk_pack rlt (plus,zero,opac)] *)
- val mk_pack: Coq.reltype -> pack -> Term.constr
-
- (** [zero rlt pack]*)
- val zero: Coq.reltype -> Term.constr -> Term.constr
-
- (** [plus rlt pack]*)
- val plus: Coq.reltype -> Term.constr -> Term.constr
-
+ val mk_pack: AAC_coq.Relation.t -> pack -> Term.constr
+
+ (** [null] builds a dummy (identity) symbol, given an {!AAC_coq.Relation.t} *)
+ val null: AAC_coq.Relation.t -> Term.constr
+
end
-(** Dynamically typed A operations *)
-module Prd:
-sig
-
- (** mimics the Coq record [Prd.pack] *)
- type pack = {dot: Term.constr; one: Term.constr ; opa: Term.constr}
-
- val typ: Term.constr lazy_t
-
- (** [mk_pack rlt (dot,one,opa)] *)
- val mk_pack: Coq.reltype -> pack -> Term.constr
-
- (** [one rlt pack]*)
- val one: Coq.reltype -> Term.constr -> Term.constr
-
- (** [dot rlt pack]*)
- val dot: Coq.reltype -> Term.constr -> Term.constr
-
+(** We need to export some Coq stubs out of this module. They are used
+ by the main tactic, see {!AAC_rewrite} *)
+module Stubs : sig
+ val lift : Term.constr Lazy.t
+ val lift_proj_equivalence : Term.constr Lazy.t
+ val lift_transitivity_left : Term.constr Lazy.t
+ val lift_transitivity_right : Term.constr Lazy.t
+ val lift_reflexivity : Term.constr Lazy.t
+ (** The evaluation fonction, used to convert a reified coq term to a
+ raw coq term *)
+ val eval: Term.constr lazy_t
+
+ (** The main lemma of our theory, that is
+ [compare (norm u) (norm v) = Eq -> eval u == eval v] *)
+ val decide_thm:Term.constr lazy_t
+
+ val lift_normalise_thm : Term.constr lazy_t
end
-
-(** [mk_equal rlt l r] builds the term [l == r] *)
-val mk_equal : Coq.reltype -> Term.constr -> Term.constr -> Term.constr
-
-
(** {2 Building reified terms}
-
+
We define a bundle of functions to build reified versions of the
terms (those that will be given to the reflexive decision
procedure). In particular, each field takes as first argument the
index of the symbol rather than the symbol itself. *)
-
+
(** Tranlations between Coq and OCaml *)
module Trans : sig
(** This module provides facilities to interpret a term with
arbitrary operators as an instance of an abstract syntax tree
- {!Matcher.Terms.t}.
+ {!AAC_matcher.Terms.t}.
For each Coq application [f x_1 ... x_n], this amounts to
deciding whether one of the partial applications [f x_1
@@ -143,77 +105,93 @@ module Trans : sig
whether the whole term is the unit for some A or AC operator. We
use typeclass resolution to test each of these possibilities.
- Note that there are ambiguous terms:
+ Note that there are ambiguous terms:
- a term like [f x y] might yield a unary morphism ([f x]) and a
binary one ([f]); we select the latter one (unless [f] is A or
AC, in which case we declare it accordingly);
- - a term like [S O] can be considered as a morphism ([S])
- applied to a unit for [(+)], or as a unit for [( * )]; we
- chose to give priority to units, so that the latter
- interpretation is selected in this case;
- - an element might be the unit for several operations; the
- behaviour in this case is almost unspecified: we simply give
- priority to AC operations.
+ - a term like [S O] can be considered as a morphism ([S])
+ applied to a unit for [(+)], or as a unit for [( * )]; we
+ chose to give priority to units, so that the latter
+ interpretation is selected in this case;
+ - an element might be the unit for several operations
*)
-
-
+
(** To achieve this reification, one need to record informations
about the collected operators (symbols, binary operators,
units). We use the following imperative internal data-structure to
this end. *)
-
- type envs
+
+ type envs
val empty_envs : unit -> envs
+
- (** {1 Reification: from Coq terms to AST {!Matcher.Terms.t} } *)
+ (** {2 Reification: from Coq terms to AST {!AAC_matcher.Terms.t} } *)
- (** [t_of_constr goal rlt envs cstr] builds the abstract syntax tree
- of the term [cstr]. We rely on the [goal] to perform typeclasses
- resolutions to find morphisms compatible with the relation
- [rlt]. Doing so, it modifies the reification environment
- [envs]. Moreover, we need to create creates fresh evars; this is
- why we give back the [goal], accordingly updated. *)
- val t_of_constr : Coq.goal_sigma -> Coq.reltype -> envs -> Term.constr -> Matcher.Terms.t * Coq.goal_sigma
- (** {1 Reconstruction: from AST back to Coq terms }
-
+ (** [t_of_constr goal rlt envs (left,right)] builds the abstract
+ syntax tree of the terms [left] and [right]. We rely on the [goal]
+ to perform typeclasses resolutions to find morphisms compatible
+ with the relation [rlt]. Doing so, it modifies the reification
+ environment [envs]. Moreover, we need to create fresh
+ evars; this is why we give back the [goal], accordingly
+ updated. *)
+
+ val t_of_constr : AAC_coq.goal_sigma -> AAC_coq.Relation.t -> envs -> (Term.constr * Term.constr) -> AAC_matcher.Terms.t * AAC_matcher.Terms.t * AAC_coq.goal_sigma
+
+ (** [add_symbol] adds a given binary symbol to the environment of
+ known stuff. *)
+ val add_symbol : AAC_coq.goal_sigma -> AAC_coq.Relation.t -> envs -> Term.constr -> AAC_coq.goal_sigma
+
+ (** {2 Reconstruction: from AST back to Coq terms }
+
The next functions allow one to map OCaml abstract syntax trees
to Coq terms. We need two functions to rebuild different kind of
terms: first, raw terms, like the one encountered by
{!t_of_constr}; second, reified Coq terms, that are required for
the reflexive decision procedure. *)
- (** {2 Building raw, natural, terms} *)
+ type ir
+ val ir_of_envs : AAC_coq.goal_sigma -> AAC_coq.Relation.t -> envs -> AAC_coq.goal_sigma * ir
+ val ir_to_units : ir -> AAC_matcher.ext_units
- (** [raw_constr_of_t] rebuilds a term in the raw representation *)
- val raw_constr_of_t : envs -> Coq.reltype -> Matcher.Terms.t -> Term.constr
+ (** {2 Building raw, natural, terms} *)
+
+ (** [raw_constr_of_t] rebuilds a term in the raw representation, and
+ reconstruct the named products on top of it. In particular, this
+ allow us to print the context put around the left (or right)
+ hand side of a pattern. *)
+ val raw_constr_of_t : ir -> AAC_coq.Relation.t -> (Term.rel_context) ->AAC_matcher.Terms.t -> Term.constr
(** {2 Building reified terms} *)
-
+
(** The reification environments, as Coq constrs *)
+
type sigmas = {
env_sym : Term.constr;
- env_sum : Term.constr;
- env_prd : Term.constr;
+ env_bin : Term.constr;
+ env_units : Term.constr; (* the [idx -> X:constr] *)
}
+
+
(** We need to reify two terms (left and right members of a goal)
that share the same reification envirnoment. Therefore, we need
to add letins to the proof context in order to ensure some
- sharing in the proof terms we produce.
+ sharing in the proof terms we produce.
Moreover, in order to have as much sharing as possible, we also
add letins for various partial applications that are used
throughout the terms.
-
+
To achieve this, we decompose the reconstruction function into
two steps: first, we build the reification environment and then
reify each term successively.*)
type reifier
- val mk_reifier : Coq.eqtype -> envs -> Coq.goal_sigma -> sigmas * reifier * Proof_type.tactic
+ val mk_reifier : AAC_coq.Relation.t -> Term.constr -> ir -> (sigmas * reifier -> Proof_type.tactic) -> Proof_type.tactic
(** [reif_constr_of_t reifier t] rebuilds the term [t] in the
reified form. *)
- val reif_constr_of_t : reifier -> Matcher.Terms.t -> Term.constr
+ val reif_constr_of_t : reifier -> AAC_matcher.Terms.t -> Term.constr
+
end
diff --git a/Caveats.v b/Caveats.v
new file mode 100644
index 0000000..8cef602
--- /dev/null
+++ b/Caveats.v
@@ -0,0 +1,377 @@
+(***************************************************************************)
+(* This is part of aac_tactics, it is distributed under the terms of the *)
+(* GNU Lesser General Public License version 3 *)
+(* (see file LICENSE for more details) *)
+(* *)
+(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
+(***************************************************************************)
+
+(** * Currently known caveats and limitations of the [aac_tactics] library.
+
+ Depending on your installation, either uncomment the following two
+ lines, or add them to your .coqrc files, replacing "."
+ with the path to the [aac_tactics] library
+*)
+
+Add Rec LoadPath "." as AAC_tactics.
+Add ML Path ".".
+Require Import AAC.
+Require Instances.
+
+(** ** Limitations *)
+
+(** *** 1. Dependent parameters
+ The type of the rewriting hypothesis must be of the form
+
+ [forall (x_1: T_1) ... (x_n: T_n), R l r],
+
+ where [R] is a relation over some type [T] and such that for all
+ variable [x_i] appearing in the left-hand side ([l]), we actually
+ have [T_i]=[T]. The goal should be of the form [S g d], where [S]
+ is a relation on [T].
+
+ In other words, we cannot instantiate arguments of an exogeneous
+ type. *)
+
+Section parameters.
+
+ Context {X} {R} {E: @Equivalence X R}
+ {plus} {plus_A: Associative R plus} {plus_C: Commutative R plus}
+ {plus_Proper: Proper (R ==> R ==> R) plus}
+ {zero} {Zero: Unit R plus zero}.
+
+ Notation "x == y" := (R x y) (at level 70).
+ Notation "x + y" := (plus x y) (at level 50, left associativity).
+ Notation "0" := (zero).
+
+ Variable f: nat -> X -> X.
+
+ (** in [Hf], the parameter [n] has type [nat], it cannot be instantiated automatically *)
+ Hypothesis Hf: forall n x, f n x + x == x.
+ Hypothesis Hf': forall n, Proper (R ==> R) (f n).
+
+ Goal forall a b k, a + f k (b+a) + b == a+b.
+ intros.
+ Fail aac_rewrite Hf. (** [aac_rewrite] does not instantiate [n] automatically *)
+ aac_rewrite (Hf k). (** of course, this argument can be given explicitly *)
+ aac_reflexivity.
+ Qed.
+
+ (** for the same reason, we cannot handle higher-order parameters (here, [g])*)
+ Hypothesis H : forall g x y, g x + g y == g (x + y).
+ Variable g : X -> X.
+ Hypothesis Hg : Proper (R ==> R) g.
+ Goal forall a b c, g a + g b + g c == g (a + b + c).
+ intros.
+ Fail aac_rewrite H.
+ do 2 aac_rewrite (H g). aac_reflexivity.
+ Qed.
+
+End parameters.
+
+
+(** *** 2. Exogeneous morphisms
+
+ We do not handle `exogeneous' morphisms: morphisms that move from
+ type [T] to some other type [T']. *)
+
+Section morphism.
+ Require Import NArith Minus.
+ Open Scope nat_scope.
+
+ (** Typically, although [N_of_nat] is a proper morphism from
+ [@eq nat] to [@eq N], we cannot rewrite under [N_of_nat] *)
+ Goal forall a b: nat, N_of_nat (a+b-(b+a)) = 0%N.
+ intros.
+ Fail aac_rewrite minus_diag.
+ Abort.
+
+
+ (* More generally, this prevents us from rewriting under
+ propositional contexts *)
+ Context {P} {HP : Proper (@eq nat ==> iff) P}.
+ Hypothesis H : P 0.
+
+ Goal forall a b, P (a + b - (b + a)).
+ intros a b.
+ Fail aac_rewrite minus_diag.
+ (** a solution is to introduce an evar to replace the part to be
+ rewritten. This tiresome process should be improved in the
+ future. Here, it can be done using eapply and the morphism. *)
+ eapply HP.
+ aac_rewrite minus_diag.
+ reflexivity.
+ exact H.
+ Qed.
+
+ Goal forall a b, a+b-(b+a) = 0 /\ b-b = 0.
+ intros.
+ (** similarly, we need to bring equations to the toplevel before
+ being able to rewrite *)
+ Fail aac_rewrite minus_diag.
+ split; aac_rewrite minus_diag; reflexivity.
+ Qed.
+
+End morphism.
+
+
+(** *** 3. Treatment of variance with inequations.
+
+ We do not take variance into account when we compute the set of
+ solutions to a matching problem modulo AC. As a consequence,
+ [aac_instances] may propose solutions for which [aac_rewrite] will
+ fail, due to the lack of adequate morphisms *)
+
+Section ineq.
+
+ Require Import ZArith.
+ Import Instances.Z.
+ Open Scope Z_scope.
+
+ Instance Zplus_incr: Proper (Zle ==> Zle ==> Zle) Zplus.
+ Proof. intros ? ? H ? ? H'. apply Zplus_le_compat; assumption. Qed.
+
+ Hypothesis H: forall x, x+x <= x.
+ Goal forall a b c, c + - (a + a) + b + b <= c.
+ intros.
+ (** this fails because the first solution is not valid ([Zopp] is not increasing), *)
+ Fail aac_rewrite H.
+ aac_instances H.
+ (** on the contrary, the second solution is valid: *)
+ aac_rewrite H at 1.
+ (** Currently, we cannot filter out such invalid solutions in an easy way;
+ this should be fixed in the future *)
+ Abort.
+
+End ineq.
+
+
+
+(** ** Caveats *)
+
+(** *** 1. Special treatment for units.
+ [S O] is considered as a unit for multiplication whenever a [Peano.mult]
+ appears in the goal. The downside is that [S x] does not match [1],
+ and [1] does not match [S(0+0)] whenever [Peano.mult] appears in
+ the goal. *)
+
+Section Peano.
+ Import Instances.Peano.
+
+ Hypothesis H : forall x, x + S x = S (x+x).
+
+ Goal 1 = 1.
+ (** ok (no multiplication around), [x] is instantiated with [O] *)
+ aacu_rewrite H.
+ Abort.
+
+ Goal 1*1 = 1.
+ (** fails since 1 is seen as a unit, not the application of the
+ morphism [S] to the constant [O] *)
+ Fail aacu_rewrite H.
+ Abort.
+
+ Hypothesis H': forall x, x+1 = 1+x.
+
+ Goal forall a, a+S(0+0) = 1+a.
+ (** ok (no multiplication around), [x] is instantiated with [a]*)
+ intro. aac_rewrite H'.
+ Abort.
+
+ Goal forall a, a*a+S(0+0) = 1+a*a.
+ (** fails: although [S(0+0)] is understood as the application of
+ the morphism [S] to the constant [O], it is not recognised
+ as the unit [S O] of multiplication *)
+ intro. Fail aac_rewrite H'.
+ Abort.
+
+ (** More generally, similar counter-intuitive behaviours can appear
+ when declaring an applied morphism as an unit. *)
+
+End Peano.
+
+
+
+(** *** 2. Existential variables.
+We implemented an algorithm for _matching_ modulo AC, not for
+_unifying_ modulo AC. As a consequence, existential variables
+appearing in a goal are considered as constants, they will not be
+instantiated. *)
+
+Section evars.
+ Require Import ZArith.
+ Import Instances.Z.
+
+ Variable P: Prop.
+ Hypothesis H: forall x y, x+y+x = x -> P.
+ Hypothesis idem: forall x, x+x = x.
+ Goal P.
+ eapply H.
+ aac_rewrite idem. (** this works: [x] is instantiated with an evar *)
+ instantiate (2 := 0).
+ symmetry. aac_reflexivity. (** this does work but there are remaining evars in the end *)
+ Abort.
+
+ Hypothesis H': forall x, 3+x = x -> P.
+ Goal P.
+ eapply H'.
+ Fail aac_rewrite idem. (** this fails since we do not instantiate evars *)
+ Abort.
+End evars.
+
+
+(** *** 3. Distinction between [aac_rewrite] and [aacu_rewrite] *)
+
+Section U.
+ Context {X} {R} {E: @Equivalence X R}
+ {dot} {dot_A: Associative R dot} {dot_Proper: Proper (R ==> R ==> R) dot}
+ {one} {One: Unit R dot one}.
+
+ Infix "==" := R (at level 70).
+ Infix "*" := dot.
+ Notation "1" := one.
+
+ (** In some situations, the [aac_rewrite] tactic allows
+ instantiations of a variable with a unit, when the variable occurs
+ directly under a function symbol: *)
+
+ Variable f : X -> X.
+ Hypothesis Hf : Proper (R ==> R) f.
+ Hypothesis dot_inv_left : forall x, f x*x == x.
+ Goal f 1 == 1.
+ aac_rewrite dot_inv_left. reflexivity.
+ Qed.
+
+ (** This behaviour seems desirable in most situations: these
+ solutions with units are less peculiar than the other ones, since
+ the unit comes from the goal. However, this policy is not properly
+ enforced for now (hard to do with the current algorithm): *)
+
+ Hypothesis dot_inv_right : forall x, x*f x == x.
+ Goal f 1 == 1.
+ Fail aac_rewrite dot_inv_right.
+ aacu_rewrite dot_inv_right. reflexivity.
+ Qed.
+
+End U.
+
+(** *** 4. Rewriting units *)
+Section V.
+ Context {X} {R} {E: @Equivalence X R}
+ {dot} {dot_A: Associative R dot} {dot_Proper: Proper (R ==> R ==> R) dot}
+ {one} {One: Unit R dot one}.
+
+ Infix "==" := R (at level 70).
+ Infix "*" := dot.
+ Notation "1" := one.
+
+ (** [aac_rewrite] uses the symbols appearing in the goal and the
+ hypothesis to infer the AC and A operations. In the following
+ example, [dot] appears neither in the left-hand-side of the goal,
+ nor in the right-hand side of the hypothesis. Hence, 1 is not
+ recognised as a unit. To circumvent this problem, we can force
+ [aac_rewrite] to take into account a given operation, by giving
+ it an extra argument. This extra argument seems useful only in
+ this peculiar case. *)
+
+ Lemma inv_unique: forall x y y', x*y == 1 -> y'*x == 1 -> y==y'.
+ Proof.
+ intros x y y' Hxy Hy'x.
+ aac_instances <- Hy'x [dot].
+ aac_rewrite <- Hy'x at 1 [dot].
+ aac_rewrite Hxy. aac_reflexivity.
+ Qed.
+End V.
+
+(** *** 5. Rewriting too much things. *)
+Section W.
+ Variables a b c: nat.
+ Hypothesis H: 0 = c.
+
+ Goal b*(a+a) <= b*(c+a+a+1).
+
+ (** [aac_rewrite] finds a pattern modulo AC that matches a given
+ hypothesis, and then makes a call to [setoid_rewrite]. This
+ [setoid_rewrite] can unfortunately make several rewrites (in a
+ non-intuitive way: below, the [1] in the right-hand side is
+ rewritten into [S c]) *)
+ aac_rewrite H.
+
+ (** To this end, we provide a companion tactic to [aac_rewrite]
+ and [aacu_rewrite], that makes the transitivity step, but not the
+ setoid_rewrite:
+
+ This allows the user to select the relevant occurrences in which
+ to rewrite. *)
+ aac_pattern H at 2. setoid_rewrite H at 1.
+ Abort.
+
+End W.
+
+(** *** 6. Rewriting nullifiable patterns. *)
+Section Z.
+
+(** If the pattern of the rewritten hypothesis does not contain "hard"
+symbols (like constants, function symbols, AC or A symbols without
+units), there can be infinitely many subterms such that the pattern
+matches: it is possible to build "subterms" modulo ACU that make the
+size of the term increase (by making neutral elements appear in a
+layered fashion). Hence, we settled with heuristics to propose only
+"some" of these solutions. In such cases, the tactic displays a
+(conservative) warning. *)
+
+Variables a b c: nat.
+Variable f: nat -> nat.
+Hypothesis H0: forall x, 0 = x - x.
+Hypothesis H1: forall x, 1 = x * x.
+
+Goal a+b*c = c.
+ aac_instances H0.
+ (** In this case, only three solutions are proposed, while there are
+ infinitely many solutions. E.g.
+ - a+b*c*(1+[])
+ - a+b*c*(1+0*(1+ []))
+ - ...
+ *)
+Abort.
+
+(** **** If the pattern is a unit:
+
+ The heuristic is to make the unit appear at each possible position
+ in the term, e.g. transforming [a] into [1*a] and [a*1], but this
+ process is not recursive (we will not transform [1*a]) into
+ [(1+0*1)*a] *)
+
+Goal a+b+c = c.
+
+ aac_instances H0 [mult].
+ (** 1 solution, we miss solutions like [(a+b+c*(1+0*(1+[])))] and so on *)
+
+ aac_instances H1 [mult].
+ (** 7 solutions, we miss solutions like [(a+b+c+0*(1+0*[]))]*)
+Abort.
+
+(** **** If the pattern is a nullifiable, but not a unit:
+
+ The heuristic is to make units appear around subterms that are
+built with AC or A symbols. Note that this case is relevant only if
+one uses [aacu_rewrite]. *)
+
+Hypothesis H : forall x y, (x+y)*x = x*x + y*x.
+Goal a*a+b*a + c = c.
+ (** Here, only one solution if we use the aac_instance tactic *)
+ aac_instances <- H.
+
+ (** There are three solutions using aacu_instances (but, here,
+ there are infinitely many different solutions). We miss e.g. [a*a +b*a
+ + (x*x + y*x)*c], which seems to be more peculiar. *)
+ aacu_instances <- H.
+Abort.
+
+
+(** The behavior of the tactic is not satisfying in this case. It is
+still unclear how to handle properly this kind of situation : we plan
+to investigate on this in the future *)
+
+End Z.
+
diff --git a/Instances.v b/Instances.v
index 1a2e08f..bb309fe 100644
--- a/Instances.v
+++ b/Instances.v
@@ -6,60 +6,176 @@
(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
(***************************************************************************)
+
+Require Export AAC.
+
(** Instances for aac_rewrite.*)
-(* At the moment, all the instances are exported even if they are packaged into modules. *)
-Require Export AAC.
+(* This one is not declared as an instance: this interferes badly with setoid_rewrite *)
+Lemma eq_subr {X} {R} `{@Reflexive X R}: subrelation eq R.
+Proof. intros x y ->. reflexivity. Qed.
+
+(* At the moment, all the instances are exported even if they are packaged into modules. Even using LocalInstances in fact*)
Module Peano.
Require Import Arith NArith Max.
- Program Instance aac_plus : @Op_AC nat eq plus 0 := @Build_Op_AC nat (@eq nat) plus 0 _ plus_0_l plus_assoc plus_comm.
+ Instance aac_plus_Assoc : Associative eq plus := plus_assoc.
+ Instance aac_plus_Comm : Commutative eq plus := plus_comm.
+
+ Instance aac_mult_Comm : Commutative eq mult := mult_comm.
+ Instance aac_mult_Assoc : Associative eq mult := mult_assoc.
+
+ Instance aac_min_Comm : Commutative eq min := min_comm.
+ Instance aac_min_Assoc : Associative eq min := min_assoc.
+
+ Instance aac_max_Comm : Commutative eq max := max_comm.
+ Instance aac_max_Assoc : Associative eq max := max_assoc.
+
+ Instance aac_one : Unit eq mult 1 := Build_Unit eq mult 1 mult_1_l mult_1_r.
+ Instance aac_zero_plus : Unit eq plus O := Build_Unit eq plus (O) plus_0_l plus_0_r.
+ Instance aac_zero_max : Unit eq max O := Build_Unit eq max 0 max_0_l max_0_r.
+
+
+ (* We also provide liftings from le to eq *)
+ Instance preorder_le : PreOrder le := Build_PreOrder _ _ le_refl le_trans.
+ Instance lift_le_eq : AAC_lift le eq := Build_AAC_lift eq_equivalence _.
-
- Program Instance aac_mult : Op_AC eq mult 1 := Build_Op_AC _ _ _ mult_assoc mult_comm.
- (* We also declare a default associative operation: this is currently
- required to fill reification environments *)
- Definition default_a := AC_A aac_plus. Existing Instance default_a.
End Peano.
+
Module Z.
- Require Import ZArith.
+ Require Import ZArith Zminmax.
Open Scope Z_scope.
- Program Instance aac_plus : Op_AC eq Zplus 0 := Build_Op_AC _ _ _ Zplus_assoc Zplus_comm.
- Program Instance aac_mult : Op_AC eq Zmult 1 := Build_Op_AC _ _ Zmult_1_l Zmult_assoc Zmult_comm.
- Definition default_a := AC_A aac_plus. Existing Instance default_a.
+ Instance aac_Zplus_Assoc : Associative eq Zplus := Zplus_assoc.
+ Instance aac_Zplus_Comm : Commutative eq Zplus := Zplus_comm.
+
+ Instance aac_Zmult_Comm : Commutative eq Zmult := Zmult_comm.
+ Instance aac_Zmult_Assoc : Associative eq Zmult := Zmult_assoc.
+
+ Instance aac_Zmin_Comm : Commutative eq Zmin := Zmin_comm.
+ Instance aac_Zmin_Assoc : Associative eq Zmin := Zmin_assoc.
+
+ Instance aac_Zmax_Comm : Commutative eq Zmax := Zmax_comm.
+ Instance aac_Zmax_Assoc : Associative eq Zmax := Zmax_assoc.
+
+ Instance aac_one : Unit eq Zmult 1 := Build_Unit eq Zmult 1 Zmult_1_l Zmult_1_r.
+ Instance aac_zero_Zplus : Unit eq Zplus 0 := Build_Unit eq Zplus 0 Zplus_0_l Zplus_0_r.
+
+ (* We also provide liftings from le to eq *)
+ Instance preorder_Zle : PreOrder Zle := Build_PreOrder _ _ Zle_refl Zle_trans.
+ Instance lift_le_eq : AAC_lift Zle eq := Build_AAC_lift eq_equivalence _.
+
End Z.
+Module Lists.
+ Require Import List.
+ Instance aac_append_Assoc {A} : Associative eq (@app A) := @app_assoc A.
+ Instance aac_nil_append {A} : @Unit (list A) eq (@app A) (@nil A) := Build_Unit _ (@app A) (@nil A) (@app_nil_l A) (@app_nil_r A).
+ Instance aac_append_Proper {A} : Proper (eq ==> eq ==> eq) (@app A).
+ Proof.
+ repeat intro.
+ subst.
+ reflexivity.
+ Qed.
+End Lists.
+
+
+Module N.
+ Require Import NArith Nminmax.
+ Open Scope N_scope.
+ Instance aac_Nplus_Assoc : Associative eq Nplus := Nplus_assoc.
+ Instance aac_Nplus_Comm : Commutative eq Nplus := Nplus_comm.
+
+ Instance aac_Nmult_Comm : Commutative eq Nmult := Nmult_comm.
+ Instance aac_Nmult_Assoc : Associative eq Nmult := Nmult_assoc.
+
+ Instance aac_Nmin_Comm : Commutative eq Nmin := N.min_comm.
+ Instance aac_Nmin_Assoc : Associative eq Nmin := N.min_assoc.
+
+ Instance aac_Nmax_Comm : Commutative eq Nmax := N.max_comm.
+ Instance aac_Nmax_Assoc : Associative eq Nmax := N.max_assoc.
+
+ Instance aac_one : Unit eq Nmult (1)%N := Build_Unit eq Nmult (1)%N Nmult_1_l Nmult_1_r.
+ Instance aac_zero : Unit eq Nplus (0)%N := Build_Unit eq Nplus (0)%N Nplus_0_l Nplus_0_r.
+ Instance aac_zero_max : Unit eq Nmax 0 := Build_Unit eq Nmax 0 N.max_0_l N.max_0_r.
+
+ (* We also provide liftings from le to eq *)
+ Instance preorder_le : PreOrder Nle := Build_PreOrder _ Nle N.T.le_refl N.T.le_trans.
+ Instance lift_le_eq : AAC_lift Nle eq := Build_AAC_lift eq_equivalence _.
+
+End N.
+
+Module P.
+ Require Import NArith Pminmax.
+ Open Scope positive_scope.
+ Instance aac_Pplus_Assoc : Associative eq Pplus := Pplus_assoc.
+ Instance aac_Pplus_Comm : Commutative eq Pplus := Pplus_comm.
+
+ Instance aac_Pmult_Comm : Commutative eq Pmult := Pmult_comm.
+ Instance aac_Pmult_Assoc : Associative eq Pmult := Pmult_assoc.
+
+ Instance aac_Pmin_Comm : Commutative eq Pmin := P.min_comm.
+ Instance aac_Pmin_Assoc : Associative eq Pmin := P.min_assoc.
+
+ Instance aac_Pmax_Comm : Commutative eq Pmax := P.max_comm.
+ Instance aac_Pmax_Assoc : Associative eq Pmax := P.max_assoc.
+
+ Instance aac_one : Unit eq Pmult 1 := Build_Unit eq Pmult 1 _ Pmult_1_r.
+ intros; reflexivity. Qed. (* TODO : add this lemma in the stdlib *)
+ Instance aac_one_max : Unit eq Pmax 1 := Build_Unit eq Pmax 1 P.max_1_l P.max_1_r.
+
+ (* We also provide liftings from le to eq *)
+ Instance preorder_le : PreOrder Ple := Build_PreOrder _ Ple P.T.le_refl P.T.le_trans.
+ Instance lift_le_eq : AAC_lift Ple eq := Build_AAC_lift eq_equivalence _.
+End P.
+
Module Q.
- Require Import QArith.
- Program Instance aac_plus : Op_AC Qeq Qplus 0 := Build_Op_AC _ _ Qplus_0_l Qplus_assoc Qplus_comm.
- Program Instance aac_mult : Op_AC Qeq Qmult 1 := Build_Op_AC _ _ Qmult_1_l Qmult_assoc Qmult_comm.
- Definition default_a := AC_A aac_plus. Existing Instance default_a.
+ Require Import QArith Qminmax.
+ Instance aac_Qplus_Assoc : Associative Qeq Qplus := Qplus_assoc.
+ Instance aac_Qplus_Comm : Commutative Qeq Qplus := Qplus_comm.
+
+ Instance aac_Qmult_Comm : Commutative Qeq Qmult := Qmult_comm.
+ Instance aac_Qmult_Assoc : Associative Qeq Qmult := Qmult_assoc.
+
+ Instance aac_Qmin_Comm : Commutative Qeq Qmin := Q.min_comm.
+ Instance aac_Qmin_Assoc : Associative Qeq Qmin := Q.min_assoc.
+
+ Instance aac_Qmax_Comm : Commutative Qeq Qmax := Q.max_comm.
+ Instance aac_Qmax_Assoc : Associative Qeq Qmax := Q.max_assoc.
+
+ Instance aac_one : Unit Qeq Qmult 1 := Build_Unit Qeq Qmult 1 Qmult_1_l Qmult_1_r.
+ Instance aac_zero_Qplus : Unit Qeq Qplus 0 := Build_Unit Qeq Qplus 0 Qplus_0_l Qplus_0_r.
+
+ (* We also provide liftings from le to eq *)
+ Instance preorder_le : PreOrder Qle := Build_PreOrder _ Qle Q.T.le_refl Q.T.le_trans.
+ Instance lift_le_eq : AAC_lift Qle Qeq := Build_AAC_lift QOrderedType.QOrder.TO.eq_equiv _.
+
End Q.
Module Prop_ops.
- Program Instance aac_or : Op_AC iff or False. Solve All Obligations using tauto.
-
- Program Instance aac_and : Op_AC iff and True. Solve All Obligations using tauto.
-
- Definition default_a := AC_A aac_and. Existing Instance default_a.
-
+ Instance aac_or_Assoc : Associative iff or. Proof. unfold Associative; tauto. Qed.
+ Instance aac_or_Comm : Commutative iff or. Proof. unfold Commutative; tauto. Qed.
+ Instance aac_and_Assoc : Associative iff and. Proof. unfold Associative; tauto. Qed.
+ Instance aac_and_Comm : Commutative iff and. Proof. unfold Commutative; tauto. Qed.
+ Instance aac_True : Unit iff or False. Proof. constructor; firstorder. Qed.
+ Instance aac_False : Unit iff and True. Proof. constructor; firstorder. Qed.
+
Program Instance aac_not_compat : Proper (iff ==> iff) not.
Solve All Obligations using firstorder.
+
+ Instance lift_impl_iff : AAC_lift Basics.impl iff := Build_AAC_lift _ _.
End Prop_ops.
Module Bool.
- Program Instance aac_orb : Op_AC eq orb false.
- Solve All Obligations using firstorder.
-
- Program Instance aac_andb : Op_AC eq andb true.
- Solve All Obligations using firstorder.
-
- Definition default_a := AC_A aac_andb. Existing Instance default_a.
-
- Instance negb_compat : Proper (eq ==> eq) negb.
- Proof. intros [|] [|]; auto. Qed.
+ Instance aac_orb_Assoc : Associative eq orb. Proof. unfold Associative; firstorder. Qed.
+ Instance aac_orb_Comm : Commutative eq orb. Proof. unfold Commutative; firstorder. Qed.
+ Instance aac_andb_Assoc : Associative eq andb. Proof. unfold Associative; firstorder. Qed.
+ Instance aac_andb_Comm : Commutative eq andb. Proof. unfold Commutative; firstorder. Qed.
+ Instance aac_true : Unit eq orb false. Proof. constructor; firstorder. Qed.
+ Instance aac_false : Unit eq andb true. Proof. constructor; intros [|];firstorder. Qed.
+
+ Instance negb_compat : Proper (eq ==> eq) negb. Proof. intros [|] [|]; auto. Qed.
End Bool.
Module Relations.
@@ -75,17 +191,20 @@ Module Relations.
Definition bot : relation T := fun _ _ => False.
Definition top : relation T := fun _ _ => True.
End defs.
-
- Program Instance aac_union T : Op_AC (same_relation T) (union T) (bot T).
- Solve All Obligations using compute; [tauto || intuition].
-
- Program Instance aac_inter T : Op_AC (same_relation T) (inter T) (top T).
- Solve All Obligations using compute; firstorder.
-
+
+ Instance eq_same_relation T : Equivalence (same_relation T). Proof. firstorder. Qed.
+
+ Instance aac_union_Comm T : Commutative (same_relation T) (union T). Proof. unfold Commutative; compute; intuition. Qed.
+ Instance aac_union_Assoc T : Associative (same_relation T) (union T). Proof. unfold Associative; compute; intuition. Qed.
+ Instance aac_bot T : Unit (same_relation T) (union T) (bot T). Proof. constructor; compute; intuition. Qed.
+
+ Instance aac_inter_Comm T : Commutative (same_relation T) (inter T). Proof. unfold Commutative; compute; intuition. Qed.
+ Instance aac_inter_Assoc T : Associative (same_relation T) (inter T). Proof. unfold Associative; compute; intuition. Qed.
+ Instance aac_top T : Unit (same_relation T) (inter T) (top T). Proof. constructor; compute; intuition. Qed.
+
(* note that we use [eq] directly as a neutral element for composition *)
- Program Instance aac_compo T : Op_A (same_relation T) (compo T) eq.
- Solve All Obligations using compute; firstorder.
- Solve All Obligations using compute; firstorder subst; trivial.
+ Instance aac_compo T : Associative (same_relation T) (compo T). Proof. unfold Associative; compute; firstorder. Qed.
+ Instance aac_eq T : Unit (same_relation T) (compo T) (eq). Proof. compute; firstorder subst; trivial. Qed.
Instance negr_compat T : Proper (same_relation T ==> same_relation T) (negr T).
Proof. compute. firstorder. Qed.
@@ -94,57 +213,43 @@ Module Relations.
Proof. compute. firstorder. Qed.
Instance clos_trans_incr T : Proper (inclusion T ==> inclusion T) (clos_trans T).
- Proof.
- intros R S H x y Hxy. induction Hxy.
+ Proof.
+ intros R S H x y Hxy. induction Hxy.
constructor 1. apply H. assumption.
- econstructor 2; eauto 3.
+ econstructor 2; eauto 3.
Qed.
- Instance clos_trans_compat T : Proper (same_relation T ==> same_relation T) (clos_trans T).
- Proof. intros R S H; split; apply clos_trans_incr, H. Qed.
+ Instance clos_trans_compat T: Proper (same_relation T ==> same_relation T) (clos_trans T).
+ Proof. intros R S H; split; apply clos_trans_incr, H. Qed.
Instance clos_refl_trans_incr T : Proper (inclusion T ==> inclusion T) (clos_refl_trans T).
- Proof.
- intros R S H x y Hxy. induction Hxy.
+ Proof.
+ intros R S H x y Hxy. induction Hxy.
constructor 1. apply H. assumption.
constructor 2.
- econstructor 3; eauto 3.
+ econstructor 3; eauto 3.
Qed.
Instance clos_refl_trans_compat T : Proper (same_relation T ==> same_relation T) (clos_refl_trans T).
- Proof. intros R S H; split; apply clos_refl_trans_incr, H. Qed.
-
+ Proof. intros R S H; split; apply clos_refl_trans_incr, H. Qed.
+
+ Instance preorder_inclusion T : PreOrder (inclusion T).
+ Proof. constructor; unfold Reflexive, Transitive, inclusion; intuition. Qed.
+
+ Instance lift_inclusion_same_relation T: AAC_lift (inclusion T) (same_relation T) :=
+ Build_AAC_lift (eq_same_relation T) _.
+ Proof. firstorder. Qed.
+
End Relations.
Module All.
Export Peano.
Export Z.
+ Export P.
+ Export N.
Export Prop_ops.
Export Bool.
Export Relations.
End All.
-
-(* A small test to ensure that everything can live together *)
-(* Section test.
- Import All.
- Require Import ZArith.
- Open Scope Z_scope.
- Notation "x ^2" := (x*x) (at level 40).
- Hypothesis H : forall x y, x^2 + y^2 + x*y + x* y = (x+y)^2.
- Goal forall a b c, a*1*(a ^2)*a + c + (b+0)*1*b + a^2*b + a*b*a = (a^2+b)^2+c.
- intros.
- aac_rewrite H.
- aac_rewrite <- H .
- symmetry.
- aac_rewrite <- H .
- aac_reflexivity.
- Qed.
- Open Scope nat_scope.
- Notation "x ^^2" := (x * x) (at level 40).
- Hypothesis H' : forall (x y : nat), x^^2 + y^^2 + x*y + x* y = (x+y)^^2.
-
- Goal forall (a b c : nat), a*1*(a ^^2)*a + c + (b+0)*1*b + a^^2*b + a*b*a = (a^^2+b)^^2+c.
- intros. aac_rewrite H'. aac_reflexivity.
- Qed.
-End test.
-
-*)
-
+
+(* Here, we should not see any instance of our classes.
+ Print HintDb typeclass_instances.
+*)
diff --git a/README.txt b/README.txt
index ad3b5c9..ae945cc 100644
--- a/README.txt
+++ b/README.txt
@@ -15,26 +15,35 @@ FOREWORD
This plugin provides tactics for rewriting universally quantified
equations, modulo associativity and commutativity of some operators.
-INSTALL
-=======
+INSTALLATION
+============
+
+This plugin should work with Coq v8.3, as released on the 14th of
+October 2010.
- run [./make_makefile] in the top-level directory to generate a Makefile
- run [make world] to build the plugin and the documentation
-This should work with Coq v8.3 `beta' : -r 13027, -r 13060 and hopefully the
-following ones.
+Option 1
+--------
+To install the plugin in Coq's directory, as, e.g., omega or ring.
+
+- run [sudo make install CMXSFILES='aac_tactics.cmxs aac_tactics.cma']
+ to copy the relevant files of the plugin
-To be able to import the plugin from anywhere, add the following to
-your ~/.coqrc
+Option 2
+--------
-Add Rec LoadPath "path_to_aac_tactics".
-Add ML Path "path_to_aac_tactics".
+If you chose not to use the previous option, you will need to add the
+following lines to (all) your .v files to be able to use the plugin:
+Add Rec LoadPath "absolute_path_to_aac_tactics".
+Add ML Path "absolute_path_to_aac_tactics".
DOCUMENTATION
=============
-Please refer to Tutorial.v for a succint introduction on how to use
+Please refer to Tutorial.v for a succinct introduction on how to use
this plugin.
To understand the inner-working of the tactic, please refer to the
@@ -48,6 +57,8 @@ we have instances for:
- Peano naturals (Import Instances.Peano)
- Z binary numbers (Import Instances.Z)
+- N binary numbers (Import Instances.N)
+- P binary numbers (Import Instances.P)
- Rationnal numbers (Import Instances.Q)
- Prop (Import Instances.Prop_ops)
- Booleans (Import Instances.Bool)
diff --git a/Tests.v b/Tests.v
deleted file mode 100644
index 87cc121..0000000
--- a/Tests.v
+++ /dev/null
@@ -1,373 +0,0 @@
-(***************************************************************************)
-(* This is part of aac_tactics, it is distributed under the terms of the *)
-(* GNU Lesser General Public License version 3 *)
-(* (see file LICENSE for more details) *)
-(* *)
-(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
-(***************************************************************************)
-
-Require Import AAC.
-
-Section sanity.
-
- Context {X} {R} {E: Equivalence R} {plus} {zero}
- {dot} {one} {A: @Op_A X R dot one} {AC: Op_AC R plus zero}.
-
-
- Notation "x == y" := (R x y) (at level 70).
- Notation "x * y" := (dot x y) (at level 40, left associativity).
- Notation "1" := (one).
- Notation "x + y" := (plus x y) (at level 50, left associativity).
- Notation "0" := (zero).
-
- Section t0.
- Hypothesis H : forall x, x+x == x.
- Goal forall a b, a+b +a == a+b.
- intros a b; aacu_rewrite H; aac_reflexivity.
- Qed.
- End t0.
-
- Section t1.
- Variables a b : X.
-
- Variable P : X -> Prop.
- Hypothesis H : forall x y, P y -> x+y+x == x.
- Hypothesis Hb : P b.
- Goal a+b +a == a.
- aacu_rewrite H.
- aac_reflexivity.
- auto.
- Qed.
- End t1.
-
- Section t.
- Variable f : X -> X.
- Hypothesis Hf : Proper (R ==> R) f.
- Hypothesis H : forall x, x+x == x.
- Goal forall y, f(y+y) == f (y).
- intros y.
- aacu_instances H.
- aacu_rewrite H.
- aac_reflexivity.
- Qed.
-
- Goal forall y z, f(y+z+z+y) == f (y+z).
- intros y z.
- aacu_instances H.
- aacu_rewrite H.
- aac_reflexivity.
- Qed.
-
- Hypothesis H' : forall x, x*x == x.
- Goal forall y z, f(y*z*y*z) == f (y*z).
- intros y z.
- aacu_instances H'.
- aacu_rewrite H'.
- aac_reflexivity.
- Qed.
- End t.
-
- Goal forall x y, (plus x y) == (plus y x). intros.
- aac_reflexivity.
- Qed.
-
- Goal forall x y, R (plus (dot x one) y) (plus y x). intros.
- aac_reflexivity.
- Qed.
-
- Goal (forall f y, f y + f y == y) -> forall a b g (Hg: Proper (R ==> R) g), g (a+b) + g b + g (b+a) == a + b + g b.
- intros.
- try aacu_rewrite H.
- aacu_rewrite (H g).
- Restart.
- intros.
- set (H' := H g).
- aacu_rewrite H'.
- aac_reflexivity.
- Qed.
-
- Section t2.
- Variables a b: X.
- Variables g : X ->X.
- Hypothesis Hg:Proper (R ==> R) g.
- Hypothesis H: (forall y, (g y) + y == y).
- Goal g (a+b) + b +a+a == a + b +a.
- intros.
- aacu_rewrite H.
- aac_reflexivity.
- Qed.
- End t2.
- Section t3.
- Variables a b: X.
- Variables g : X ->X.
- Hypothesis Hg:Proper (R ==> R) g.
- Hypothesis H: (forall f y, f y + y == y).
-
- Goal g (a+b) + b +a+a == a + b +a.
- intros.
- aacu_rewrite (H g).
- aac_reflexivity.
- Qed.
-End t3.
-
-Section null.
- Hypothesis dot_ann : forall x, 0 * x == 0.
- Goal forall a, 0*a == 0.
- intros a. aac_rewrite dot_ann. reflexivity.
- Qed.
-End null.
-
-End sanity.
-
-Section abs.
- Context
- {X} {E: Equivalence (@eq X)}
- {plus} {zero} {AC: @Op_AC X eq plus zero}
- {plus'} {zero'} {AC': @Op_AC X eq plus' zero'}
- {dot} {one} {A: @Op_A X eq dot one}
- {dot'} {one'} {A': @Op_A X eq dot' one'}.
-
- Notation "x * y" := (dot x y) (at level 40, left associativity).
- Notation "x @ y" := (dot' x y) (at level 20, left associativity).
- Notation "1" := (one).
- Notation "1'" := (one').
- Notation "x + y" := (plus x y) (at level 50, left associativity).
- Notation "x ^ y" := (plus' x y) (at level 30, right associativity).
- Notation "0" := (zero).
- Notation "0'" := (zero').
-
- Variables a b c d e k: X.
- Variables f g: X -> X.
- Hypothesis Hf: Proper (eq ==> eq) f.
- Hypothesis Hg: Proper (eq ==> eq) g.
-
- Variables h: X -> X -> X.
- Hypothesis Hh: Proper (eq ==> eq ==> eq) h.
-
- Variables q: X -> X -> X -> X.
- Hypothesis Hq: Proper (eq ==> eq ==> eq ==> eq) q.
-
-
-
-
- Hypothesis dot_ann_left: forall x, 0*x = 0.
-
- Goal f (a*0*b*c*d) = k.
- aac_rewrite dot_ann_left.
- Restart.
- aac_rewrite dot_ann_left subterm 1 subst 0.
- Abort.
-
-
-
-
- Hypothesis distr_dp: forall x y z, x*(y+z) = x*y+x*z.
- Hypothesis distr_dp': forall x y z, x*y+x*z = x*(y+z).
-
- Hypothesis distr_pp: forall x y z, x^(y+z) = x^y+x^z.
- Hypothesis distr_pp': forall x y z, x^y+x^z = x^(y+z).
-
- Hypothesis distr_dd: forall x y z, x@(y*z) = x@y * x@z.
- Hypothesis distr_dd': forall x y z, x@y * x@z = x@(y*z).
-
-
- Goal a*b*(c+d+e) = k.
- aac_instances distr_dp.
- aacu_instances distr_dp.
- aac_rewrite distr_dp subterm 0 subst 4.
-
- aac_instances distr_dp'.
- aac_rewrite distr_dp'.
- Abort.
-
- Goal a@(b*c)@d@(e*e) = k.
- aacu_instances distr_dd.
- aac_instances distr_dd.
- aac_rewrite distr_dd.
- Abort.
-
- Goal a^(b+c)^d^(e+e) = k.
- aacu_instances distr_dd.
- aac_instances distr_dd.
- aacu_instances distr_pp.
- aac_instances distr_pp.
- aac_rewrite distr_pp subterm 9.
- Abort.
-
-
-
-
- Hypothesis plus_idem: forall x, x+x = x.
- Hypothesis plus_idem3: forall x, x+x+x = x.
- Hypothesis dot_idem: forall x, x*x = x.
- Hypothesis dot_idem3: forall x, x*x*x = x.
-
- Goal a*b*a*b = k.
- aac_instances dot_idem.
- aacu_instances dot_idem.
- aac_rewrite dot_idem.
- aac_instances distr_dp. (* TODO: no solution -> fail ? *)
- aacu_instances distr_dp.
- Abort.
-
- Goal a+a+a+a = k.
- aac_instances plus_idem3.
- aac_rewrite plus_idem3.
- Abort.
-
- Goal a*a*a*a = k.
- aac_instances dot_idem3.
- aac_rewrite dot_idem3.
- Abort.
-
- Goal a*a*a*a*a*a = k.
- aac_instances dot_idem3.
- aac_rewrite dot_idem3.
- Abort.
-
- Goal f(a*a)*f(a*a) = k.
- aac_instances dot_idem.
- (* TODO: pas clair qu'on doive réécrire les deux petits
- sous-termes quand on nous demande de n'en réécrire qu'un...
- d'autant plus que le comportement est nécessairement
- imprévisible (cf. les deux exemples en dessous)
- ceci dit, je suis conscient que c'est relou à empêcher... *)
- aac_rewrite dot_idem subterm 1.
- Abort.
-
- Goal f(a*a*b)*f(a*a*b) = k.
- aac_instances dot_idem.
- aac_rewrite dot_idem subterm 2. (* une seule instance réécrite *)
- Abort.
-
- Goal f(b*a*a)*f(b*a*a) = k.
- aac_instances dot_idem.
- aac_rewrite dot_idem subterm 2. (* deux instances réécrites *)
- Abort.
-
- Goal h (a*a) (a*b*c*d*e) = k.
- aac_rewrite dot_idem.
- Abort.
-
- Goal h (a+a) (a+b+c+d+e) = k.
- aac_rewrite plus_idem.
- Abort.
-
- Goal (a*a+a*a+b)*c = k.
- aac_rewrite plus_idem.
- Restart.
- aac_rewrite dot_idem.
- Abort.
-
-
-
-
- Hypothesis dot_inv: forall x, x*f x = 1.
- Hypothesis plus_inv: forall x, x+f x = 0.
- Hypothesis dot_inv': forall x, f x*x = 1.
-
- Goal a*f(1)*b = k.
- try aac_rewrite dot_inv. (* TODO: à terme, j'imagine qu'on souhaite accepter ça, même en aac *)
- aacu_rewrite dot_inv.
- Abort.
-
- Goal f(1) = k.
- aacu_rewrite dot_inv.
- Restart.
- aacu_rewrite dot_inv'.
- Abort.
-
- Goal b*f(b) = k.
- aac_rewrite dot_inv.
- Abort.
-
- Goal f(b)*b = k.
- aac_rewrite dot_inv'.
- Abort.
-
- Goal a*f(a*b)*a*b*c = k.
- aac_rewrite dot_inv'.
- Abort.
-
- Goal g (a*f(a*b)*a*b*c+d) = k.
- aac_rewrite dot_inv'.
- Abort.
-
- Goal g (a+f(a+b)+c+b+b) = k.
- aac_rewrite plus_inv.
- Abort.
-
- Goal g (a+f(0)+c) = k.
- aac_rewrite plus_inv.
- Abort.
-
-
-
-
-
- Goal forall R S, Equivalence R -> Equivalence S -> R a k -> S a k.
- intros R S HR HS H.
- try (aac_rewrite H ; fail 1 "ne doit pas passer").
- try (aac_instances H; fail 1 "ne doit pas passer").
- Abort.
-
- Goal forall R S, Equivalence R -> Equivalence S -> (forall x, R (f x) k) -> S (f a) k.
- intros R S HR HS H.
- try (aac_instances H; fail 1 "ne doit pas passer").
- try (aac_rewrite H ; fail 1 "ne doit pas passer").
- Abort.
-
-
-
-
-
- Hypothesis star_f: forall x y, x*f(y*x) = f(x*y)*x.
-
- Goal g (a+b*c*d*f(a*b*c*d)+b) = k.
- aac_instances star_f.
- aac_rewrite star_f.
- Abort.
-
- Goal b*f(a*b)*f(b*b*f(a*b)) = k.
- aac_instances star_f.
- aac_rewrite star_f.
- aac_rewrite star_f.
- Abort.
-
-
-
-
-
- Hypothesis dot_select: forall x y, f x * y * g x = g y*x*f y.
-
- Goal f(a+b)*c*g(b+a) = k.
- aac_rewrite dot_select.
- Abort.
-
- Goal f(a+b)*g(b+a) = k.
- try (aac_rewrite dot_select; fail 1 "ne doit pas passer").
- aacu_rewrite dot_select.
- Abort.
-
- Goal f b*f(a+b)*g(b+a)*g b = k.
- aacu_instances dot_select.
- aac_instances dot_select.
- aac_rewrite dot_select.
- aacu_rewrite dot_select.
- aacu_rewrite dot_select.
- Abort.
-
-
-
-
- Hypothesis galois_dot: forall x y z, h x (y*z) = h (x*y) z.
- Hypothesis galois_plus: forall x y z, h x (y+z) = h (x+y) z.
-
- Hypothesis select_dot: forall x y, h x (y*x) = h (x*y) x.
- Hypothesis select_plus: forall x y, h x (y+x) = h (x+y) x.
-
- Hypothesis h_dot: forall x y, h (x*y) (y*x) = h x y.
- Hypothesis h_plus: forall x y, h (x+y) (y+x) = h x y.
-
-End abs.
-
diff --git a/Tutorial.v b/Tutorial.v
index 84d40fc..8b4c753 100644
--- a/Tutorial.v
+++ b/Tutorial.v
@@ -6,97 +6,70 @@
(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
(***************************************************************************)
-(** * Examples about uses of the aac_rewrite library *)
+(** * Tutorial for using the [aac_tactics] library.
-(*
- Depending on your installation, either uncomment the following two
- lines, or add them to your .coqrc files, replacing path_to_aac_tactics
- with the correct path for your installation:
+ Depending on your installation, either modify the following two
+ lines, or add them to your .coqrc files, replacing "." with the
+ path to the [aac_tactics] library. *)
- Add Rec LoadPath "path_to_aac_tactics".
- Add ML Path "path_to_aac_tactics".
-*)
-Require Export AAC.
+Add Rec LoadPath "." as AAC_tactics.
+Add ML Path ".".
+Require Import AAC.
Require Instances.
-(** ** Introductory examples *)
+(** ** Introductory example
-(** First, we settle in the context of Z, and show an usage of our
-tactics: we rewrite an universally quantified hypothesis modulo
-associativity and commutativity. *)
+ Here is a first example with relative numbers ([Z]): one can
+ rewrite an universally quantified hypothesis modulo the
+ associativity and commutativity of [Zplus]. *)
Section introduction.
+
Import ZArith.
Import Instances.Z.
-
+
Variables a b c : Z.
Hypothesis H: forall x, x + Zopp x = 0.
Goal a + b + c + Zopp (c + a) = b.
aac_rewrite H.
aac_reflexivity.
Qed.
- Goal a + c+ Zopp (b + a + Zopp b) = c.
+ Goal a + c + Zopp (b + a + Zopp b) = c.
do 2 aac_rewrite H.
reflexivity.
Qed.
- (** Notes:
+ (** Notes:
- the tactic handles arbitrary function symbols like [Zopp] (as
- long as they are proper morphisms);
+ long as they are proper morphisms w.r.t. the considered
+ equivalence relation);
- here, ring would have done the job.
*)
-End introduction.
-
-(** Second, we show how to exploit binomial identities to prove a goal
-about pythagorean triples, without breaking a sweat. By comparison,
-even if the goal and the hypothesis are both in normal form, making
-the rewrites using standard tools is difficult.
-*)
-
-Section binomials.
-
- Import ZArith.
- Import Instances.Z.
-
- Notation "x ^2" := (x*x) (at level 40).
- Notation "2 â‹… x" := (x+x) (at level 41).
- Lemma Hbin1: forall x y, (x+y)^2 = x^2 + y^2 + 2â‹…x*y. Proof. intros; ring. Qed.
- Lemma Hbin2: forall x y, x^2 + y^2 = (x+y)^2 + -(2â‹…x*y). Proof. intros; ring. Qed.
- Lemma Hopp : forall x, x + -x = 0. Proof Zplus_opp_r.
-
- Variables a b c : Z.
- Hypothesis H : c^2 + 2â‹…(a+1)*b = (a+1+b)^2.
- Goal a^2 + b^2 + 2â‹…a + 1 = c^2.
- aacu_rewrite <- Hbin1.
- aac_rewrite Hbin2.
- aac_rewrite <- H.
- aac_rewrite Hopp.
- aac_reflexivity.
- Qed.
-
- (** Note: after the [aac_rewrite <- H], one could use [ring] to close the proof.*)
-
-End binomials.
+End introduction.
-(** ** Usage *)
-(** One can also work in an abstract context, with arbitrary
- associative and commutative operators.
+(** ** Usage
- (Note that one can declare several operations of each kind;
- however, to be able to use this plugin, one currently needs at least
- one associative operator, and one associative-commutative
- operator.) *)
+ One can also work in an abstract context, with arbitrary
+ associative and commutative operators. (Note that one can declare
+ several operations of each kind.) *)
Section base.
-
- Context {X} {R} {E: Equivalence R}
- {plus} {zero}
- {dot} {one}
- {A: @Op_A X R dot one}
- {AC: Op_AC R plus zero}.
-
+ Context {X} {R} {E: Equivalence R}
+ {plus}
+ {dot}
+ {zero}
+ {one}
+ {dot_A: @Associative X R dot }
+ {plus_A: @Associative X R plus }
+ {plus_C: @Commutative X R plus }
+ {dot_Proper :Proper (R ==> R ==> R) dot}
+ {plus_Proper :Proper (R ==> R ==> R) plus}
+ {Zero : Unit R plus zero}
+ {One : Unit R dot one}
+ .
+
Notation "x == y" := (R x y) (at level 70).
Notation "x * y" := (dot x y) (at level 40, left associativity).
Notation "1" := (one).
@@ -110,30 +83,32 @@ Section base.
Section reminder.
Hypothesis H : forall x, x * x == x.
Variables a b c : X.
-
+
Goal (a+b+c)*(c+a+b) == a+b+c.
aac_rewrite H.
aac_reflexivity.
Qed.
- (** Note: the tactic starts by normalizing terms, so that trailing
- units are always eliminated. *)
+ (** The tactic starts by normalising terms, so that trailing units
+ are always eliminated. *)
Goal ((a+b)+0+c)*((c+a)+b*1) == a+b+c.
aac_rewrite H.
aac_reflexivity.
Qed.
End reminder.
-
- (** We can deal with "proper" morphisms of arbitrary arity (here [f],
- or [Zopp] earlier), and rewrite under morphisms (here [g]). *)
+
+ (** The tactic can deal with "proper" morphisms of arbitrary arity
+ (here [f] and [g], or [Zopp] earlier): it rewrites under such
+ morphisms ([g]), and, more importantly, it is able to reorder
+ terms modulo AC under these morphisms ([f]). *)
Section morphisms.
Variable f : X -> X -> X.
Hypothesis Hf : Proper (R ==> R ==> R) f.
Variable g : X -> X.
Hypothesis Hg : Proper (R ==> R) g.
-
+
Variable a b: X.
Hypothesis H : forall x y, x+f (b+y) x == y+x.
Goal g ((f (a+b) a) + a) == g (a+a).
@@ -142,180 +117,283 @@ Section base.
Qed.
End morphisms.
- (** ** Selecting what and where to rewrite *)
+ (** *** Selecting what and where to rewrite
- (** There are sometimes several possible rewriting. We now show how
- to interact with the tactic to select the desired one. *)
+ There are sometimes several solutions to the matching problem. We
+ now show how to interact with the tactic to select the desired
+ one. *)
Section occurrence.
- Variable f : X -> X .
+ Variable f : X -> X.
Variable a : X.
Hypothesis Hf : Proper (R ==> R) f.
- Hypothesis H : forall x, x + x == x.
+ Hypothesis H : forall x, x + x == x.
Goal f(a+a)+f(a+a) == f a.
(** In case there are several possible solutions, one can print
the different solutions using the [aac_instances] tactic (in
- proofgeneral, look at buffer *coq* ): *)
+ proof-general, look at buffer *coq* ): *)
aac_instances H.
- (** the default choice is the smallest possible context (number
- 0), but one can choose the desired context; *)
- aac_rewrite H subterm 1.
- (** now the goal is [f a + f a == f a], there is only one solution. *)
+ (** the default choice is the occurrence with the smallest
+ possible context (number 0), but one can choose the desired
+ one; *)
+ aac_rewrite H at 1.
+ (** now the goal is [f a + f a == f a], there is only one solution. *)
aac_rewrite H.
reflexivity.
Qed.
-
+
End occurrence.
Section subst.
Variables a b c d : X.
Hypothesis H: forall x y, a*x*y*b == a*(x+y)*b.
Hypothesis H': forall x, x + x == x.
-
+
Goal a*c*d*c*d*b == a*c*d*b.
- (** Here, there is only one possible context, but several substitutions; *)
+ (** Here, there is only one possible occurrence, but several substitutions; *)
aac_instances H.
- (** we can select them with the proper keyword. *)
- aac_rewrite H subst 1.
+ (** one can select them with the proper keyword. *)
+ aac_rewrite H subst 1.
aac_rewrite H'.
aac_reflexivity.
Qed.
End subst.
-
- (** As expected, one can use both keyword together to select the
- correct subterm and the correct substitution. *)
+
+ (** As expected, one can use both keywords together to select the
+ occurrence and the substitution. We also provide a keyword to
+ specify that the rewrite should be done in the right-hand side of
+ the equation. *)
Section both.
Variables a b c d : X.
Hypothesis H: forall x y, a*x*y*b == a*(x+y)*b.
Hypothesis H': forall x, x + x == x.
-
- Goal a*c*d*c*d*b*b == a*(c*d+b)*b.
+
+ Goal a*c*d*c*d*b*b == a*(c*d*b)*b.
+ aac_instances H.
+ aac_rewrite H at 1 subst 1.
aac_instances H.
- aac_rewrite H subterm 1 subst 1.
aac_rewrite H.
aac_rewrite H'.
+ aac_rewrite H at 0 subst 1 in_right.
aac_reflexivity.
Qed.
-
+
End both.
- (** We now turn on explaining the distinction between [aac_rewrite]
- and [aacu_rewrite]: [aac_rewrite] rejects solutions in which
- variables are instantiated by units, the companion tactic,
- [aacu_rewrite] allows such solutions. *)
+ (** *** Distinction between [aac_rewrite] and [aacu_rewrite]:
+
+ [aac_rewrite] rejects solutions in which variables are instantiated
+ by units, while the companion tactic, [aacu_rewrite] allows such
+ solutions. *)
Section dealing_with_units.
Variables a b c: X.
Hypothesis H: forall x, a*x*a == x.
Goal a*a == 1.
- (** Here, x must be instantiated with 1, hence no solutions; *)
- aac_instances H.
- (** while we get solutions with the "aacu" tactic. *)
+ (** Here, [x] must be instantiated with [1], so that the [aac_*]
+ tactics give no solutions; *)
+ try aac_instances H.
+ (** while we get solutions with the [aacu_*] tactics. *)
aacu_instances H.
aacu_rewrite H.
reflexivity.
Qed.
-
+
(** We introduced this distinction because it allows us to rule
out dummy cases in common situations: *)
Hypothesis H': forall x y z, x*y + x*z == x*(y+z).
- Goal a*b*c + a*c+ a*b == a*(c+b*(1+c)).
+ Goal a*b*c + a*c + a*b == a*(c+b*(1+c)).
(** 6 solutions without units, *)
- aac_instances H'.
+ aac_instances H'.
+ aac_rewrite H' at 0.
(** more than 52 with units. *)
aacu_instances H'.
- Abort.
-
- End dealing_with_units.
+ Abort.
+
+ End dealing_with_units.
End base.
-(** ** Declaring instances *)
+(** *** Declaring instances
-(** One can use one's own operations: it suffices to declare them as
- instances of our classes. (Note that these instances are already
- declared in file [Instances.v].) *)
+ To use one's own operations: it suffices to declare them as
+ instances of our classes. (Note that the following instances are
+ already declared in file [Instances.v].) *)
Section Peano.
Require Import Arith.
-
- Program Instance nat_plus : Op_AC eq plus O.
- Solve All Obligations using firstorder.
+
+ Instance aac_plus_Assoc : Associative eq plus := plus_assoc.
+ Instance aac_plus_Comm : Commutative eq plus := plus_comm.
+
+ Instance aac_one : Unit eq mult 1 :=
+ Build_Unit eq mult 1 mult_1_l mult_1_r.
+ Instance aac_zero_plus : Unit eq plus O :=
+ Build_Unit eq plus (O) plus_0_l plus_0_r.
+
+
+ (** Two (or more) operations may share the same units: in the
+ following example, [0] is understood as the unit of [max] as well as
+ the unit of [plus]. *)
+
+ Instance aac_max_Comm : Commutative eq Max.max := Max.max_comm.
+ Instance aac_max_Assoc : Associative eq Max.max := Max.max_assoc.
- Program Instance nat_dot : Op_AC eq mult 1.
- Solve All Obligations using firstorder.
-
+ Instance aac_zero_max : Unit eq Max.max O :=
+ Build_Unit eq Max.max 0 Max.max_0_l Max.max_0_r.
+
+ Variable a b c : nat.
+ Goal Max.max (a + 0) 0 = a.
+ aac_reflexivity.
+ Qed.
+
+ (** Furthermore, several operators can be mixed: *)
+
+ Hypothesis H : forall x y z, Max.max (x + y) (x + z) = x+ Max.max y z.
+
+ Goal Max.max (a + b) (c + (a * 1)) = Max.max c b + a.
+ aac_instances H. aac_rewrite H. aac_reflexivity.
+ Qed.
+ Goal Max.max (a + b) (c + Max.max (a*1+0) 0) = a + Max.max b c.
+ aac_instances H. aac_rewrite H. aac_reflexivity.
+ Qed.
+
+
+ (** *** Working with inequations
+
+ To be able to use the tactics, the goal must be a relation [R]
+ applied to two arguments, and the rewritten hypothesis must end
+ with a relation [Q] applied to two arguments. These relations are
+ not necessarily equivalences, but they should be related
+ according to the occurrence where the rewrite takes place; we
+ leave this check to the underlying call to [setoid_rewrite]. *)
+
+ (** One can rewrite equations in the left member of inequations, *)
+ Goal (forall x, x + x = x) -> a + b + b + a <= a + b.
+ intro Hx.
+ aac_rewrite Hx.
+ reflexivity.
+ Qed.
+
+ (** or in the right member of inequations, using the [in_right] keyword *)
+ Goal (forall x, x + x = x) -> a + b <= a + b + b + a.
+ intro Hx.
+ aac_rewrite Hx in_right.
+ reflexivity.
+ Qed.
+
+ (** Similarly, one can rewrite inequations in inequations, *)
+ Goal (forall x, x + x <= x) -> a + b + b + a <= a + b.
+ intro Hx.
+ aac_rewrite Hx.
+ reflexivity.
+ Qed.
+
+ (** possibly in the right-hand side. *)
+ Goal (forall x, x <= x + x) -> a + b <= a + b + b + a.
+ intro Hx.
+ aac_rewrite <- Hx in_right.
+ reflexivity.
+ Qed.
+
+ (** [aac_reflexivity] deals with "trivial" inequations too *)
+ Goal Max.max (a + b) (c + a) <= Max.max (b + a) (c + 1*a).
+ aac_reflexivity.
+ Qed.
+
+ (** In the last three examples, there were no equivalence relation
+ involved in the goal. However, we actually had to guess the
+ equivalence relation with respect to which the operators
+ ([plus,max,0]) were AC. In this case, it was Leibniz equality
+ [eq] so that it was automatically inferred; more generally, one
+ can specify which equivalence relation to use by declaring
+ instances of the [AAC_lift] type class: *)
+
+ Instance lift_le_eq : AAC_lift le eq.
+ (** (This instance is automatically inferred because [eq] is always a
+ valid candidate, here for [le].) *)
- (** Caveat: we need at least an instance of an operator that is AC
- and one that is A for a given relation. However, one can reuse an
- AC operator as an A operator. *)
- Definition default_a := AC_A nat_dot. Existing Instance default_a.
End Peano.
-(** ** Caveats *)
-Section Peano'.
- Import Instances.Peano.
+(** *** Normalising goals
- (** 1. We have a special treatment for units, thus, [S x + x] does not
- match [S 0], which is considered as a unit (one) of the mult
- operation. *)
+ We also provide a tactic to normalise terms modulo AC. This
+ normalisation is the one we use internally. *)
- Section caveat_one.
- Definition double_plus_one x := 2*x + 1.
- Hypothesis H : forall x, S x + x = double_plus_one x.
- Goal S O = double_plus_one O.
- try aac_rewrite H. (** 0 solutions (normal since it would use 0 to instantiate x) *)
- try aacu_rewrite H. (** 0 solutions (abnormal) *)
- Abort.
- End caveat_one.
+Section AAC_normalise.
+
+ Import Instances.Z.
+ Require Import ZArith.
+ Open Scope Z_scope.
+
+ Variable a b c d : Z.
+ Goal a + (b + c*c*d) + a + 0 + d*1 = a.
+ aac_normalise.
+ Abort.
- (** 2. We cannot at the moment have two classes with the same
- units: in the following example, [0] is understood as the unit of
- [max] rather than as the unit of [plus]. *)
+End AAC_normalise.
- Section max.
- Program Instance aac_max : Op_AC eq Max.max O := Build_Op_AC _ _ _ Max.max_assoc Max.max_comm.
- Variable a : nat.
- Goal 0 + a = a.
- try aac_reflexivity.
- Abort.
- End max.
-End Peano'.
-(** 3. If some computations are possible in the goal or in the
-hypothesis, the inference mechanism we use will make the
-conversion. However, it seems that in most cases, these conversions
-can be done by hand using simpl rather than rewrite. *)
-Section Z.
- (* The order of the following lines is important in order to get the right scope afterwards. *)
- Import ZArith.
+(** ** Examples from the web page *)
+Section Examples.
+
+ Import Instances.Z.
+ Require Import ZArith.
Open Scope Z_scope.
- Opaque Zmult.
-
- Hypothesis dot_ann_left :forall x, x * 0 = 0.
- Hypothesis dot_ann_right :forall x, 0 * x = 0.
- Goal forall a, a*0 = 0.
- intros. aacu_rewrite dot_ann_left. reflexivity.
+ (** *** Reverse triangle inequality *)
+
+ Lemma Zabs_triangle : forall x y, Zabs (x + y) <= Zabs x + Zabs y .
+ Proof Zabs_triangle.
+
+ Lemma Zplus_opp_r : forall x, x + -x = 0.
+ Proof Zplus_opp_r.
+
+ (** The following morphisms are required to perform the required rewrites *)
+ Instance Zminus_compat : Proper (Zge ==> Zle) Zopp.
+ Proof. intros x y. omega. Qed.
+
+ Instance Proper_Zplus : Proper (Zle ==> Zle ==> Zle) Zplus.
+ Proof. firstorder. Qed.
+
+ Goal forall a b, Zabs a - Zabs b <= Zabs (a - b).
+ intros. unfold Zminus.
+ aac_instances <- (Zminus_diag b).
+ aac_rewrite <- (Zminus_diag b) at 3.
+ unfold Zminus.
+ aac_rewrite Zabs_triangle.
+ aac_rewrite Zplus_opp_r.
+ aac_reflexivity.
Qed.
- (** Here the tactic fails, since the 0*a is converted to 0, and no
- rewrite can occur (even though Zmult is opaque). *)
- Goal forall a, 0*a = 0.
- intros. try aacu_rewrite dot_ann_left.
- Abort.
-
- (** Here the tactic fails, since the 0*x is converted to 0 in the hypothesis. *)
- Goal forall a, a*0 = 0.
- intros. try aacu_rewrite dot_ann_right.
- Abort.
-
+
+ (** *** Pythagorean triples *)
+
+ Notation "x ^2" := (x*x) (at level 40).
+ Notation "2 â‹… x" := (x+x) (at level 41).
+
+ Lemma Hbin1: forall x y, (x+y)^2 = x^2 + y^2 + 2â‹…x*y. Proof. intros; ring. Qed.
+ Lemma Hbin2: forall x y, x^2 + y^2 = (x+y)^2 + -(2â‹…x*y). Proof. intros; ring. Qed.
+ Lemma Hopp : forall x, x + -x = 0. Proof Zplus_opp_r.
-End Z.
+ Variables a b c : Z.
+ Hypothesis H : c^2 + 2â‹…(a+1)*b = (a+1+b)^2.
+ Goal a^2 + b^2 + 2â‹…a + 1 = c^2.
+ aacu_rewrite <- Hbin1.
+ rewrite Hbin2.
+ aac_rewrite <- H.
+ aac_rewrite Hopp.
+ aac_reflexivity.
+ Qed.
+
+ (** Note: after the [aac_rewrite <- H], one could use [ring] to close the proof.*)
+
+End Examples.
+
diff --git a/aac_rewrite.ml b/aac_rewrite.ml
deleted file mode 100644
index c4c88e0..0000000
--- a/aac_rewrite.ml
+++ /dev/null
@@ -1,308 +0,0 @@
-(***************************************************************************)
-(* This is part of aac_tactics, it is distributed under the terms of the *)
-(* GNU Lesser General Public License version 3 *)
-(* (see file LICENSE for more details) *)
-(* *)
-(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
-(***************************************************************************)
-
-(** aac_rewrite -- rewriting modulo *)
-
-(* functions to handle the failures of our tactic. Some should be
- reported [anomaly], some are on behalf of the user [user_error]*)
-let anomaly msg =
- Util.anomaly ("aac_tactics: " ^ msg)
-
-let user_error msg =
- Util.error ("aac_tactics: " ^ msg)
-
-(* debug and timing functions *)
-let debug = false
-let printing = false
-let time = false
-
-let debug x =
- if debug then
- Printf.printf "%s\n%!" x
-
-let pr_constr msg constr =
- if printing then
- ( Pp.msgnl (Pp.str (Printf.sprintf "=====%s====" msg));
- Pp.msgnl (Printer.pr_constr constr);
- )
-
-let time msg tac =
- if time then Coq.tclTIME msg tac
- else tac
-
-let tac_or_exn tac exn msg =
- fun gl -> try tac gl with e -> exn msg e
-
-(* helper to be used with the previous function: raise a new anomaly
- except if a another one was previously raised *)
-let push_anomaly msg = function
- | Util.Anomaly _ as e -> raise e
- | _ -> anomaly msg
-
-module M = Matcher
-
-open Term
-open Names
-open Coqlib
-open Proof_type
-
-
-(** [find_applied_equivalence goal eq] try to ensure that the goal is
- an applied equivalence relation, with two operands of the same type.
-
- This function is transitionnal, as we plan to integrate with the
- new rewrite features of Coq 8.3, that seems to handle this kind of
- things.
-*)
-let find_applied_equivalence goal : Term.constr -> Coq.eqtype *Term.constr *Term.constr* Coq.goal_sigma = fun equation ->
- let env = Tacmach.pf_env goal in
- let evar_map = Tacmach.project goal in
- match Coq.decomp_term equation with
- | App(c,ca) when Array.length ca >= 2 ->
- let n = Array.length ca in
- let left = ca.(n-2) in
- let right = ca.(n-1) in
- let left_type = Tacmach.pf_type_of goal left in
- (* let right_type = Tacmach.pf_type_of goal right in *)
- let partial_app = mkApp (c, Array.sub ca 0 (n-2) ) in
- let eq_type = Coq.Classes.mk_equivalence left_type partial_app in
- begin
- try
- let evar_map, equivalence = Typeclasses.resolve_one_typeclass env evar_map eq_type in
- let goal = Coq.goal_update goal evar_map in
- ((left_type, partial_app), equivalence),left, right, goal
- with
- | Not_found -> user_error "The goal is not an applied equivalence"
- end
- | _ -> user_error "The goal is not an equation"
-
-
-(* [find_applied_equivalence_force] brutally decomposes the given
- equation, and fail if we find a relation that is not the same as
- the one we look for [r] *)
-let find_applied_equivalence_force rlt goal : Term.constr -> Term.constr *Term.constr* Coq.goal_sigma = fun equation ->
- match Coq.decomp_term equation with
- | App(c,ca) when Array.length ca >= 2 ->
- let n = Array.length ca in
- let left = ca.(n-2) in
- let right = ca.(n-1) in
- let partial_app = mkApp (c, Array.sub ca 0 (n-2) ) in
- let _ = if snd rlt = partial_app then () else user_error "The goal and the hypothesis have different top-level relations" in
-
- left, right, goal
- | _ -> user_error "The hypothesis is not an equation"
-
-
-
-(** Build a couple of [t] from an hypothesis (variable names are not
- relevant) *)
-let rec t_of_hyp goal rlt envs e = match Coq.decomp_term e with
- | Prod (name,ty,c) -> let x, y = t_of_hyp goal rlt envs c
- in x,y+1
- | _ ->
- let l,r ,goal = find_applied_equivalence_force rlt goal e in
- let l,goal = Theory.Trans.t_of_constr goal rlt envs l in
- let r,goal = Theory.Trans.t_of_constr goal rlt envs r in
- (l,r),0
-
-(** @deprecated: [option_rip] will be removed as soon as a proper interface is
- given to the user *)
-let option_rip = function | Some x -> x | None -> raise Not_found
-
-let mk_hyp_app conv c env =
- let l = Matcher.Subst.to_list env in
- let l = List.sort (fun (x,_) (y,_) -> compare y x) l in
- let l = List.map (fun x -> conv (snd x)) l in
- let v = Array.of_list l in
- mkApp (c,v)
-
-
-
-(** {1 Tactics} *)
-
-(** [by_aac_reflexivity] is a sub-tactic that closes a sub-goal that
- is merely a proof of equality of two terms modulo AAC *)
-
-let by_aac_reflexivity eqt envs (t : Matcher.Terms.t) (t' : Matcher.Terms.t) : Proof_type.tactic = fun goal ->
- let maps, reifier,reif_letins = Theory.Trans.mk_reifier eqt envs goal in
- let ((x,r),e) = eqt in
-
- (* fold through a term and reify *)
- let t = Theory.Trans.reif_constr_of_t reifier t in
- let t' = Theory.Trans.reif_constr_of_t reifier t' in
-
- (* Some letins *)
- let eval, eval_letin = Coq.mk_letin' "eval_name" (mkApp (Lazy.force Theory.eval, [|x;r; maps.Theory.Trans.env_sym; maps.Theory.Trans.env_prd; maps.Theory.Trans.env_sum|])) in
- let _ = pr_constr "left" t in
- let _ = pr_constr "right" t' in
- let t , t_letin = Coq.mk_letin "left" t in
- let t', t_letin'= Coq.mk_letin "right" t' in
- let apply_tac = Tactics.apply (
- mkApp (Lazy.force Theory.decide_thm, [|x;r;e; maps.Theory.Trans.env_sym; maps.Theory.Trans.env_prd; maps.Theory.Trans.env_sum;t;t'|]))
- in
- Tactics.tclABSTRACT None
- (Tacticals.tclTHENLIST
- [
- tac_or_exn reif_letins anomaly "invalid letin (1)";
- tac_or_exn eval_letin anomaly "invalid letin (2)";
- tac_or_exn t_letin anomaly "invalid letin (3)";
- tac_or_exn t_letin' anomaly "invalid letin (4)";
- tac_or_exn apply_tac anomaly "could not apply the decision theorem";
- tac_or_exn (time "vm_norm" (Tactics.normalise_vm_in_concl)) anomaly "vm_compute failure";
- Tacticals.tclORELSE Tactics.reflexivity
- (Tacticals.tclFAIL 0 (Pp.str "Not an equality modulo A/AC"))
- ])
- goal
-
-
-(** The core tactic for aac_rewrite *)
-let core_aac_rewrite ?(l2r = true) envs eqt rew p subst left = fun goal ->
- let rlt = fst eqt in
- let t = Matcher.Subst.instantiate subst p in
- let tr = Theory.Trans.raw_constr_of_t envs rlt t in
- let tac_rew_l2r =
- if l2r then Equality.rewriteLR rew else Equality.rewriteRL rew
- in
- pr_constr "transitivity through" tr;
- Tacticals.tclTHENSV
- (tac_or_exn (Tactics.transitivity tr) anomaly "Unable to make the transitivity step")
- [|
- (* si by_aac_reflexivity foire (pas un theoreme), il fait tclFAIL, et on rattrape ici *)
- tac_or_exn (time "by_aac_refl" (by_aac_reflexivity eqt envs left t)) push_anomaly "Invalid transitivity step";
- tac_or_exn (tac_rew_l2r) anomaly "Unable to rewrite";
- |] goal
-
-exception NoSolutions
-
-(** Choose a substitution from a
- [(int * Terms.t * Env.env Search.m) Search.m ] *)
-let choose_subst subterm sol m=
- try
- let (depth,pat,envm) = match subterm with
- | None -> (* first solution *)
- option_rip (Matcher.Search.choose m)
- | Some x ->
- List.nth (List.rev (Matcher.Search.to_list m)) x
- in
- let env = match sol with
- None -> option_rip (Matcher.Search.choose envm)
- | Some x -> List.nth (List.rev (Matcher.Search.to_list envm)) x
- in
- pat, env
- with
- | _ -> raise NoSolutions
-
-(** rewrite the constr modulo AC from left to right in the
- left member of the goal
-*)
-let aac_rewrite rew ?(l2r=true) ?(show = false) ?strict ?occ_subterm ?occ_sol : Proof_type.tactic = fun goal ->
- let envs = Theory.Trans.empty_envs () in
- let rew_type = Tacmach.pf_type_of goal rew in
- let (gl : Term.types) = Tacmach.pf_concl goal in
- let _ = pr_constr "h" rew_type in
- let _ = pr_constr "gl" gl in
- let eqt, left, right, goal = find_applied_equivalence goal gl in
- let rlt = fst eqt in
- let left,goal = Theory.Trans.t_of_constr goal rlt envs left in
- let right,goal = Theory.Trans.t_of_constr goal rlt envs right in
- let (hleft,hright),hn = t_of_hyp goal rlt envs rew_type in
- let pattern = if l2r then hleft else hright in
- let m = Matcher.subterm ?strict pattern left in
-
- let m = Matcher.Search.sort (fun (x,_,_) (y,_,_) -> x - y) m in
- if show
- then
- let _ = Pp.msgnl (Pp.str "All solutions:") in
- let _ = Pp.msgnl (Matcher.pp_all (fun t -> Printer.pr_constr (Theory.Trans.raw_constr_of_t envs rlt t) ) m) in
- Tacticals.tclIDTAC goal
- else
- try
- let pat,env = choose_subst occ_subterm occ_sol m in
- let rew = mk_hyp_app (Theory.Trans.raw_constr_of_t envs rlt )rew env in
- core_aac_rewrite ~l2r envs eqt rew pat env left goal
- with
- | NoSolutions -> Tacticals.tclFAIL 0
- (Pp.str (if occ_subterm = None && occ_sol = None
- then "No matching subterm found"
- else "No such solution")) goal
-
-
-let aac_reflexivity : Proof_type.tactic = fun goal ->
- let (gl : Term.types) = Tacmach.pf_concl goal in
- let envs = Theory.Trans.empty_envs() in
- let eqt, left, right, evar_map = find_applied_equivalence goal gl in
- let rlt = fst eqt in
- let left,goal = Theory.Trans.t_of_constr goal rlt envs left in
- let right,goal = Theory.Trans.t_of_constr goal rlt envs right in
- by_aac_reflexivity eqt envs left right goal
-
-
-open Tacmach
-open Tacticals
-open Tacexpr
-open Tacinterp
-open Extraargs
-
- (* Adding entries to the grammar of tactics *)
-TACTIC EXTEND _aac1_
-| [ "aac_reflexivity" ] -> [ aac_reflexivity ]
-END
-
-
-TACTIC EXTEND _aac2_
-| [ "aac_rewrite" orient(o) constr(c) "subterm" integer(i) "subst" integer(j)] ->
- [ fun gl -> aac_rewrite ~l2r:o ~strict:true ~occ_subterm:i ~occ_sol:j c gl ]
-END
-
-TACTIC EXTEND _aac3_
-| [ "aac_rewrite" orient(o) constr(c) "subst" integer(j)] ->
- [ fun gl -> aac_rewrite ~l2r:o ~strict:true ~occ_subterm:0 ~occ_sol:j c gl ]
-END
-
-TACTIC EXTEND _aac4_
-| [ "aac_rewrite" orient(o) constr(c) "subterm" integer(i)] ->
- [ fun gl -> aac_rewrite ~l2r:o ~strict:true ~occ_subterm:i ~occ_sol:0 c gl ]
-END
-
-TACTIC EXTEND _aac5_
-| [ "aac_rewrite" orient(o) constr(c) ] ->
- [ fun gl -> aac_rewrite ~l2r:o ~strict:true c gl ]
-END
-
-TACTIC EXTEND _aac6_
-| [ "aac_instances" orient(o) constr(c)] ->
- [ fun gl -> aac_rewrite ~l2r:o ~strict:true ~show:true c gl ]
-END
-
-
-
-
-TACTIC EXTEND _aacu2_
-| [ "aacu_rewrite" orient(o) constr(c) "subterm" integer(i) "subst" integer(j)] ->
- [ fun gl -> aac_rewrite ~l2r:o ~occ_subterm:i ~occ_sol:j c gl ]
-END
-
-TACTIC EXTEND _aacu3_
-| [ "aacu_rewrite" orient(o) constr(c) "subst" integer(j)] ->
- [ fun gl -> aac_rewrite ~l2r:o ~occ_subterm:0 ~occ_sol:j c gl ]
-END
-
-TACTIC EXTEND _aacu4_
-| [ "aacu_rewrite" orient(o) constr(c) "subterm" integer(i)] ->
- [ fun gl -> aac_rewrite ~l2r:o ~occ_subterm:i ~occ_sol:0 c gl ]
-END
-
-TACTIC EXTEND _aacu5_
-| [ "aacu_rewrite" orient(o) constr(c) ] ->
- [ fun gl -> aac_rewrite ~l2r:o c gl ]
-END
-
-TACTIC EXTEND _aacu6_
-| [ "aacu_instances" orient(o) constr(c)] ->
- [ fun gl -> aac_rewrite ~l2r:o ~show:true c gl ]
-END
diff --git a/aac_rewrite.mli b/aac_rewrite.mli
deleted file mode 100644
index 81818b6..0000000
--- a/aac_rewrite.mli
+++ /dev/null
@@ -1,59 +0,0 @@
-(***************************************************************************)
-(* This is part of aac_tactics, it is distributed under the terms of the *)
-(* GNU Lesser General Public License version 3 *)
-(* (see file LICENSE for more details) *)
-(* *)
-(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
-(***************************************************************************)
-
-(** Main module for the Coq plug-in ; provides the [aac_rewrite] and
- [aac_reflexivity] tactics.
-
- This file defines the entry point for the tactic aac_rewrite. It
- does Joe-the-plumbing, given a goal, reifies the interesting
- subterms (rewrited hypothesis, goal) into abstract syntax tree
- {!Matcher.Terms} (see {!Theory.Trans.t_of_constr}). Then, we use the
- results from the matcher to rebuild terms and make a transitivity
- step toward a term in which the hypothesis can be rewritten using
- the standard rewrite.
-
- Doing so, we generate a sub-goal which we solve using a reflexive
- decision procedure for the equality of terms modulo
- AAC. Therefore, we also need to reflect the goal into a concrete
- data-structure. See {i AAC.v} for more informations,
- especially the data-type {b T} and the {b decide} theorem.
-
-*)
-
-(** {2 Transitional functions}
-
- We define some plumbing functions that will be removed when we
- integrate the new rewrite features of Coq 8.3
-*)
-
-(** [find_applied_equivalence goal eq] checks that the goal is
- an applied equivalence relation, with two operands of the same
- type.
-
-*)
-val find_applied_equivalence : Proof_type.goal Tacmach.sigma -> Term.constr -> Coq.eqtype * Term.constr * Term.constr * Proof_type.goal Tacmach.sigma
-
-(** Build a couple of [t] from an hypothesis (variable names are not
- relevant) *)
-val t_of_hyp : Proof_type.goal Tacmach.sigma -> Coq.reltype -> Theory.Trans.envs -> Term.types -> (Matcher.Terms.t * Matcher.Terms.t) * int
-
-(** {2 Tactics} *)
-
-(** the [aac_reflexivity] tactic solves equalities modulo AAC, by
- reflection: it reifies the goal to apply theorem [decide], from
- file {i AAC.v}, and then concludes using [vm_compute]
- and [reflexivity]
-*)
-val aac_reflexivity : Proof_type.tactic
-
-(** [aac_rewrite] is the tactic for in-depth reqwriting modulo AAC
-with some options to choose the orientation of the rewriting and a
-solution (first the subterm, then the solution)*)
-
-val aac_rewrite : Term.constr -> ?l2r:bool -> ?show:bool -> ?strict: bool -> ?occ_subterm:int -> ?occ_sol:int -> Proof_type.tactic
-
diff --git a/coq.ml b/coq.ml
deleted file mode 100644
index 731204c..0000000
--- a/coq.ml
+++ /dev/null
@@ -1,212 +0,0 @@
-(***************************************************************************)
-(* This is part of aac_tactics, it is distributed under the terms of the *)
-(* GNU Lesser General Public License version 3 *)
-(* (see file LICENSE for more details) *)
-(* *)
-(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
-(***************************************************************************)
-
-(** Interface with Coq *)
-
-type constr = Term.constr
-
-open Term
-open Names
-open Coqlib
-
-
-(* The contrib name is used to locate errors when loading constrs *)
-let contrib_name = "aac_tactics"
-
-(* Getting constrs (primitive Coq terms) from exisiting Coq
- libraries. *)
-let find_constant contrib dir s =
- Libnames.constr_of_global (Coqlib.find_reference contrib dir s)
-
-let init_constant dir s = find_constant contrib_name dir s
-
-(* A clause specifying that the [let] should not try to fold anything
- in the goal *)
-let nowhere =
- { Tacexpr.onhyps = Some [];
- Tacexpr.concl_occs = Rawterm.no_occurrences_expr
- }
-
-
-let mk_letin (name: string) (constr: constr) : constr * Proof_type.tactic =
- let name = (Names.id_of_string name) in
- let letin =
- (Tactics.letin_tac
- None
- (Name name)
- constr
- None (* or Some ty *)
- nowhere
- ) in
- mkVar name, letin
-
-let mk_letin' (name: string) (constr: constr) : constr * Proof_type.tactic
- =
- constr, Tacticals.tclIDTAC
-
-
-
-(** {2 General functions} *)
-
-type goal_sigma = Proof_type.goal Tacmach.sigma
-let goal_update (goal : goal_sigma) evar_map : goal_sigma=
- let it = Tacmach.sig_it goal in
- Tacmach.re_sig it evar_map
-
-let resolve_one_typeclass goal ty : constr*goal_sigma=
- let env = Tacmach.pf_env goal in
- let evar_map = Tacmach.project goal in
- let em,c = Typeclasses.resolve_one_typeclass env evar_map ty in
- c, (goal_update goal em)
-
-let nf_evar goal c : Term.constr=
- let evar_map = Tacmach.project goal in
- Evarutil.nf_evar evar_map c
-
-let evar_unit (gl : goal_sigma) (rlt : constr * constr) : constr * goal_sigma =
- let env = Tacmach.pf_env gl in
- let evar_map = Tacmach.project gl in
- let (em,x) = Evarutil.new_evar evar_map env (fst rlt) in
- x,(goal_update gl em)
-
-let evar_binary (gl: goal_sigma) (rlt: constr * constr) =
- let env = Tacmach.pf_env gl in
- let evar_map = Tacmach.project gl in
- let x = fst rlt in
- let ty = mkArrow x (mkArrow x x) in
- let (em,x) = Evarutil.new_evar evar_map env ty in
- x,( goal_update gl em)
-
-(* decomp_term : constr -> (constr, types) kind_of_term *)
-let decomp_term c = kind_of_term (strip_outer_cast c)
-
-
-(** {2 Bindings with Coq' Standard Library} *)
-
-module Pair = struct
-
- let path = ["Coq"; "Init"; "Datatypes"]
- let typ = lazy (init_constant path "prod")
- let pair = lazy (init_constant path "pair")
-end
-
-module Pos = struct
-
- let path = ["Coq" ; "NArith"; "BinPos"]
- let typ = lazy (init_constant path "positive")
- let xI = lazy (init_constant path "xI")
- let xO = lazy (init_constant path "xO")
- let xH = lazy (init_constant path "xH")
-
- (* A coq positive from an int *)
- let of_int n =
- let rec aux n =
- begin match n with
- | n when n < 1 -> assert false
- | 1 -> Lazy.force xH
- | n -> mkApp
- (
- (if n mod 2 = 0
- then Lazy.force xO
- else Lazy.force xI
- ), [| aux (n/2)|]
- )
- end
- in
- aux n
-end
-
-module Nat = struct
- let path = ["Coq" ; "Init"; "Datatypes"]
- let typ = lazy (init_constant path "nat")
- let _S = lazy (init_constant path "S")
- let _O = lazy (init_constant path "O")
- (* A coq nat from an int *)
- let of_int n =
- let rec aux n =
- begin match n with
- | n when n < 0 -> assert false
- | 0 -> Lazy.force _O
- | n -> mkApp
- (
- (Lazy.force _S
- ), [| aux (n-1)|]
- )
- end
- in
- aux n
-end
-
-(** Lists from the standard library*)
-module List = struct
- let path = ["Coq"; "Lists"; "List"]
- let typ = lazy (init_constant path "list")
- let nil = lazy (init_constant path "nil")
- let cons = lazy (init_constant path "cons")
- let cons ty h t =
- mkApp (Lazy.force cons, [| ty; h ; t |])
- let nil ty =
- (mkApp (Lazy.force nil, [| ty |]))
- let rec of_list ty = function
- | [] -> nil ty
- | t::q -> cons ty t (of_list ty q)
- let type_of_list ty =
- mkApp (Lazy.force typ, [|ty|])
-end
-
-(** Morphisms *)
-module Classes =
-struct
- let classes_path = ["Coq";"Classes"; ]
- let morphism = lazy (init_constant (classes_path@["Morphisms"]) "Proper")
-
- (** build the constr [Proper rel t] *)
- let mk_morphism ty rel t =
- mkApp (Lazy.force morphism, [| ty; rel; t|])
-
- let equivalence = lazy (init_constant (classes_path@ ["RelationClasses"]) "Equivalence")
-
- (** build the constr [Equivalence ty rel] *)
- let mk_equivalence ty rel =
- mkApp (Lazy.force equivalence, [| ty; rel|])
-end
-
-(** a [mset] is a ('a * pos) list *)
-let mk_mset ty (l : (constr * int) list) =
- let pos = Lazy.force Pos.typ in
- let pair_ty = mkApp (Lazy.force Pair.typ , [| ty; pos|]) in
- let pair (x : constr) (ar : int) =
- mkApp (Lazy.force Pair.pair, [| ty ; pos ; x ; Pos.of_int ar|]) in
- let rec aux = function
- | [] -> List.nil pair_ty
- | (t,ar)::q -> List.cons pair_ty (pair t ar) (aux q)
- in
- aux l
-
-(** x r *)
-type reltype = Term.constr * Term.constr
-(** x r e *)
-type eqtype = reltype * Term.constr
-
-
-(** {1 Tacticals} *)
-
-let tclTIME msg tac = fun gl ->
- let u = Sys.time () in
- let r = tac gl in
- let _ = Pp.msgnl (Pp.str (Printf.sprintf "%s:%fs" msg (Sys.time ()-. u))) in
- r
-
-let tclDEBUG msg tac = fun gl ->
- let _ = Pp.msgnl (Pp.str msg) in
- tac gl
-
-let tclPRINT tac = fun gl ->
- let _ = Pp.msgnl (Printer.pr_constr (Tacmach.pf_concl gl)) in
- tac gl
-
diff --git a/coq.mli b/coq.mli
deleted file mode 100644
index 588a922..0000000
--- a/coq.mli
+++ /dev/null
@@ -1,117 +0,0 @@
-(***************************************************************************)
-(* This is part of aac_tactics, it is distributed under the terms of the *)
-(* GNU Lesser General Public License version 3 *)
-(* (see file LICENSE for more details) *)
-(* *)
-(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
-(***************************************************************************)
-
-(** Interface with Coq where we import several definitions from Coq's
- standard library.
-
- This general purpose library could be reused by other plugins.
-
- Some salient points:
- - we use Caml's module system to mimic Coq's one, and avoid cluttering
- the namespace;
- - we also provide several handlers for standard coq tactics, in
- order to provide a unified setting (we replace functions that
- modify the evar_map by functions that operate on the whole
- goal, and repack the modified evar_map with it).
-
-*)
-
-(** {2 Getting Coq terms from the environment} *)
-
-val init_constant : string list -> string -> Term.constr
-
-(** {2 General purpose functions} *)
-
-type goal_sigma = Proof_type.goal Tacmach.sigma
-val goal_update : goal_sigma -> Evd.evar_map -> goal_sigma
-val resolve_one_typeclass : Proof_type.goal Tacmach.sigma -> Term.types -> Term.constr * goal_sigma
-val nf_evar : goal_sigma -> Term.constr -> Term.constr
-val evar_unit :goal_sigma ->Term.constr*Term.constr-> Term.constr* goal_sigma
-val evar_binary: goal_sigma -> Term.constr*Term.constr -> Term.constr* goal_sigma
-
-
-(** [mk_letin name v] binds the constr [v] using a letin tactic *)
-val mk_letin : string ->Term.constr ->Term.constr * Proof_type.tactic
-
-(** [mk_letin' name v] is a drop-in replacement for [mk_letin' name v]
- that does not make a binding (useful to test whether using lets is
- efficient) *)
-val mk_letin' : string ->Term.constr ->Term.constr * Proof_type.tactic
-
-(* decomp_term : constr -> (constr, types) kind_of_term *)
-val decomp_term : Term.constr -> (Term.constr , Term.types) Term.kind_of_term
-
-
-(** {2 Bindings with Coq' Standard Library} *)
-
-(** Coq lists *)
-module List:
-sig
- (** [of_list ty l] *)
- val of_list:Term.constr ->Term.constr list ->Term.constr
-
- (** [type_of_list ty] *)
- val type_of_list:Term.constr ->Term.constr
-end
-
-(** Coq pairs *)
-module Pair:
-sig
- val typ:Term.constr lazy_t
- val pair:Term.constr lazy_t
-end
-
-(** Coq positive numbers (pos) *)
-module Pos:
-sig
- val typ:Term.constr lazy_t
- val of_int: int ->Term.constr
-end
-
-(** Coq unary numbers (peano) *)
-module Nat:
-sig
- val typ:Term.constr lazy_t
- val of_int: int ->Term.constr
-end
-
-
-(** Coq typeclasses *)
-module Classes:
-sig
- val mk_morphism: Term.constr -> Term.constr -> Term.constr -> Term.constr
- val mk_equivalence: Term.constr -> Term.constr -> Term.constr
-end
-
-(** Both in OCaml and Coq, we represent finite multisets using
- weighted lists ([('a*int) list]), see {!Matcher.mset}.
-
- [mk_mset ty l] constructs a Coq multiset from an OCaml multiset
- [l] of Coq terms of type [ty] *)
-
-val mk_mset:Term.constr -> (Term.constr * int) list ->Term.constr
-
-(** pairs [(x,r)] such that [r: Relation x] *)
-type reltype = Term.constr * Term.constr
-
-(** triples [((x,r),e)] such that [e : @Equivalence x r]*)
-type eqtype = reltype * Term.constr
-
-
-(** {2 Some tacticials} *)
-
-(** time the execution of a tactic *)
-val tclTIME : string -> Proof_type.tactic -> Proof_type.tactic
-
-(** emit debug messages to see which tactics are failing *)
-val tclDEBUG : string -> Proof_type.tactic -> Proof_type.tactic
-
-(** print the current goal *)
-val tclPRINT : Proof_type.tactic -> Proof_type.tactic
-
-
diff --git a/files.txt b/files.txt
index 7fa1801..fb0b175 100644
--- a/files.txt
+++ b/files.txt
@@ -1,7 +1,11 @@
-matcher.ml
-coq.ml
-theory.ml
-aac_rewrite.ml
+AAC_coq.ml
+AAC_helper.ml
+AAC_search_monad.ml
+AAC_matcher.ml
+AAC_theory.ml
+AAC_print.ml
+AAC_rewrite.ml
AAC.v
Instances.v
Tutorial.v
+Caveats.v
diff --git a/magic.txt b/magic.txt
index d0272c3..d1c9191 100644
--- a/magic.txt
+++ b/magic.txt
@@ -1,4 +1,5 @@
-custom "mkdir -p doc; $(CAMLBIN)ocamldoc -html -rectypes -d doc -m A $(ZDEBUG) $(ZFLAGS) $^ && touch doc" "$(MLIFILES)" doc
-custom "$(CAMLOPTLINK) $(ZDEBUG) $(ZFLAGS) -shared -linkall -o aac_tactics.cmxs $(CMXFILES)" "$(CMXFILES)" aac_tactics.cmxs
+-custom "$(CAMLOPTLINK) -g -a -o aac_tactics.cmxa $(CMXFILES)" "$(CMXFILES)" aac_tactics.cmxa
-custom "$(CAMLLINK) -g -a -o aac_tactics.cma $(CMOFILES)" "$(CMOFILES)" aac_tactics.cma
-custom "$(CAMLBIN)ocamldep -slash $(OCAMLLIBS) $^ > .depend" "$(MLIFILES)" .depend
diff --git a/make_makefile b/make_makefile
index 3cfc096..1cef3d3 100755
--- a/make_makefile
+++ b/make_makefile
@@ -1,35 +1,48 @@
+#!/bin/sh
# - set variable MLIFILES so that it can be used in custom targets to
# build documentation and dependencies for .mli files
# - remove useless ' -I . .../plugins/ ' lines so that we can read something in the compilation log
# - rename the rule for Makefile auto-regeneration, and replace it with an appropriate command
# - patch the rule for cleaning doc
# - include the .depend custom target, to take .mli dependencies into account
+# - added some hackish settings of Variables, in order to be able to install the plugin in user-contrib (coq_makefile should define a "INSTALL" variables that we can modify, instead of copying 4 times the same code in install !)
+
+set -e
+
+if [ -f `ocamlc -where`/dynlink.cmxa ]; then
+ BEST=opt
+ CMXA=aac_tactics.cmxa
+ CMXS=aac_tactics.cmxs
+else
+ BEST=byte
+ if which ocamlopt >/dev/null; then
+ CMXA=aac_tactics.cmxa
+ fi
+fi
(
-coq_makefile -no-install -opt MLIFILES = '$(MLFILES:.ml=.mli)' $(cat files.txt) -f magic.txt \
- | sed 's|.*/plugins/.*||g;s|Makefile:|\.dummy:|g;s|rm -f doc|rm -rf doc|g'
+coq_makefile -R . AAC_tactics -$BEST MLIFILES = '$(MLFILES:.ml=.mli)' $(cat files.txt) -f magic.txt \
+ | sed 's|.*/plugins/.*||g;s|Makefile:|\.dummy:|g;s|rm -f doc|rm -rf doc|g;s|\.opt||g'
cat <<EOF
-# sanity checks
-tests: Tests.vo
# override inadequate coq_makefile auto-regeneration
Makefile: make_makefile magic.txt files.txt
- . make_makefile
+ ./make_makefile
-# We want to keep the proofs only in the Tutorial
-tutorial: Tutorial.vo Tutorial.glob
- \$(COQDOC) --no-index --no-externals -html \$(COQDOCLIBS) -d html Tutorial.v
-
-world: \$(VOFILES) doc gallinahtml tutorial
+# We want to keep the proofs in Caveats and the Tutorial
+world: \$(VOFILES) doc
+ - mkdir -p html
+ \$(COQDOC) -toc -utf8 -html -g \$(COQDOCLIBS) -d html \$(VFILES)
+ \$(COQDOC) --no-index --no-externals -s -utf8 -html \$(COQDOCLIBS) -d html Tutorial.v
+ \$(COQDOC) --no-index --no-externals -s -utf8 -html \$(COQDOCLIBS) -d html Caveats.v
# additional dependencies for AAC (since aac_tactics.cmxs/cma are not
# around the first time we run make, coqdep issues a warning and
# ignores this dependency)
# (one can safely select one of theses dependencies according to
# coq' compilation mode--bytecode or native)
-AAC.vo: aac_tactics.cmxs aac_tactics.cma
-
+AAC.vo: $CMXA $CMXS aac_tactics.cma
# .depend contains dependencies for .mli files
-include .depend
diff --git a/matcher.ml b/matcher.ml
deleted file mode 100644
index 176b660..0000000
--- a/matcher.ml
+++ /dev/null
@@ -1,980 +0,0 @@
-(***************************************************************************)
-(* This is part of aac_tactics, it is distributed under the terms of the *)
-(* GNU Lesser General Public License version 3 *)
-(* (see file LICENSE for more details) *)
-(* *)
-(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
-(***************************************************************************)
-
-(** This module defines our matching functions, modulo associativity
- and commutativity (AAC).
-
- The basic idea is to find a substitution [env] such that the
- pattern [p] instantiated by [env] is equal to [t] modulo AAC.
-
- We proceed by structural decomposition of the pattern, and try all
- possible non-deterministic split of the subject when needed. The
- function {!matcher} is limited to top-level matching, that is, the
- subject must make a perfect match against the pattern ([x+x] do
- not match [a+a+b] ). We use a search monad {!Search} to perform
- non-deterministic splits in an almost transparent way. We also
- provide a function {!subterm} for finding a match that is a
- subterm modulo AAC of the subject. Therefore, we are able to solve
- the aforementioned case [x+x] against [a+b+a].
-
- This file is structured as follows. First, we define the search
- monad. Then,we define the two representations of terms (one
- representing the AST, and one in normal form ), and environments
- from variables to terms. Then, we use these parts to solve
- matching problem. Finally, we wrap this function {!matcher} into
- {!subterm}
-*)
-
-
-let debug = false
-let time = false
-
-
-let time f x fmt =
- if time then
- let t = Sys.time() in
- let r = f x in
- Printf.printf fmt (Sys.time () -. t);
- r
- else f x
-
-
-
-type symbol = int
-type var = int
-
-
-(****************)
-(* Search Monad *)
-(****************)
-
-
-(** The {!Search} module contains a search monad that allows to
- express, in a legible maner, programs that solves combinatorial
- problems
-
- @see <http://spivey.oriel.ox.ac.uk/mike/search-jfp.pdf> the
- inspiration of this module
-*)
-module Search : sig
- (** A data type that represent a collection of ['a] *)
- type 'a m
- (** bind and return *)
- val ( >> ) : 'a m -> ('a -> 'b m) -> 'b m
- val return : 'a -> 'a m
- (** non-deterministic choice *)
- val ( >>| ) : 'a m -> 'a m -> 'a m
- (** failure *)
- val fail : unit -> 'a m
- (** folding through the collection *)
- val fold : ('a -> 'b -> 'b) -> 'a m -> 'b -> 'b
- (** derived facilities *)
- val sprint : ('a -> string) -> 'a m -> string
- val count : 'a m -> int
- val choose : 'a m -> 'a option
- val to_list : 'a m -> 'a list
- val sort : ('a -> 'a -> int) -> 'a m -> 'a m
- val is_empty: 'a m -> bool
-end
-= struct
-
- type 'a m = | F of 'a
- | N of 'a m list
-
- let fold (f : 'a -> 'b -> 'b) (m : 'a m) (acc : 'b) =
- let rec aux acc = function
- F x -> f x acc
- | N l ->
- (List.fold_left (fun acc x ->
- match x with
- | (N []) -> acc
- | x -> aux acc x
- ) acc l)
- in
- aux acc m
-
-
-
- let rec (>>) : 'a m -> ('a -> 'b m) -> 'b m =
- fun m f ->
- match m with
- | F x -> f x
- | N l ->
- N (List.fold_left (fun acc x ->
- match x with
- | (N []) -> acc
- | x -> (x >> f)::acc
- ) [] l)
-
- let (>>|) (m : 'a m) (n :'a m) : 'a m = match (m,n) with
- | N [],_ -> n
- | _,N [] -> m
- | F x, N l -> N (F x::l)
- | N l, F x -> N (F x::l)
- | x,y -> N [x;y]
-
- let return : 'a -> 'a m = fun x -> F x
- let fail : unit -> 'a m = fun () -> N []
-
- let sprint f m =
- fold (fun x acc -> Printf.sprintf "%s\n%s" acc (f x)) m ""
- let rec count = function
- | F _ -> 1
- | N l -> List.fold_left (fun acc x -> acc+count x) 0 l
-
- let opt_comb f x y = match x with None -> f y | _ -> x
-
- let rec choose = function
- | F x -> Some x
- | N l -> List.fold_left (fun acc x ->
- opt_comb choose acc x
- ) None l
-
- let is_empty = fun x -> choose x = None
-
- let to_list m = (fold (fun x acc -> x::acc) m [])
-
- let sort f m =
- N (List.map (fun x -> F x) (List.sort f (to_list m)))
-end
-
-open Search
-
-
-type 'a mset = ('a * int) list
-let linear p =
- let rec ncons t l = function
- | 0 -> l
- | n -> t::ncons t l (n-1)
- in
- let rec aux = function
- [ ] -> []
- | (t,n)::q -> let q = aux q in
- ncons t q n
- in aux p
-
-
-
-(** The module {!Terms} defines two different types for expressions.
-
- - a public type {!Terms.t} that represent abstract syntax trees of
- expressions with binary associative (and commutative) operators
-
- - a private type {!Terms.nf_term} that represent an equivalence
- class for terms that are equal modulo AAC. The constructions
- functions on this type ensure the property that the term is in
- normal form (that is, no sum can appear as a subterm of the same
- sum, no trailing units, etc...).
-
-*)
-
-module Terms : sig
-
- (** {1 Abstract syntax tree of terms}
-
- Terms represented using this datatype are representation of the
- AST of an expression. *)
-
- type t =
- Dot of (symbol * t * t)
- | One of symbol
- | Plus of (symbol * t * t)
- | Zero of symbol
- | Sym of (symbol * t array)
- | Var of var
-
- val equal_aac : t -> t -> bool
- val size: t -> int
- (** {1 Terms in normal form}
-
- A term in normal form is the canonical representative of the
- equivalence class of all the terms that are equal modulo
- Associativity and Commutativity. Outside the {!Matcher} module,
- one does not need to access the actual representation of this
- type. *)
-
- type nf_term = private
- | TAC of symbol * nf_term mset
- | TA of symbol * nf_term list
- | TSym of symbol *nf_term list
- | TVar of var
-
-
- (** {2 Constructors: we ensure that the terms are always
- normalised} *)
- val mk_TAC : symbol -> nf_term mset -> nf_term
- val mk_TA : symbol -> nf_term list -> nf_term
- val mk_TSym : symbol -> nf_term list -> nf_term
- val mk_TVar : var -> nf_term
-
- (** {2 Comparisons} *)
-
- val nf_term_compare : nf_term -> nf_term -> int
- val nf_equal : nf_term -> nf_term -> bool
-
- (** {2 Printing function} *)
- val sprint_nf_term : nf_term -> string
-
- (** {2 Conversion functions} *)
- val term_of_t : t -> nf_term
- val t_of_term : nf_term -> t
-end
- = struct
-
- type t =
- Dot of (symbol * t * t)
- | One of symbol
- | Plus of (symbol * t * t)
- | Zero of symbol
- | Sym of (symbol * t array)
- | Var of var
-
- let rec size = function
- | Dot (_,x,y)
- | Plus (_,x,y) -> size x+ size y + 1
- | Sym (_,v)-> Array.fold_left (fun acc x -> size x + acc) 1 v
- | _ -> 1
-
-
- type nf_term =
- | TAC of symbol * nf_term mset
- | TA of symbol * nf_term list
- | TSym of symbol *nf_term list
- | TVar of var
-
-
-
- (** {2 Comparison} *)
-
- let nf_term_compare = Pervasives.compare
- let nf_equal a b = a = b
-
- (** {2 Constructors: we ensure that the terms are always
- normalised} *)
-
- (** {3 Pre constructors : These constructors ensure that sums and
- products are not degenerated (no trailing units)} *)
- let mk_TAC' s l = match l with
- | [t,0] -> TAC (s,[])
- | [t,1] -> t
- | _ -> TAC (s,l)
- let mk_TA' s l = match l with [t] -> t
- | _ -> TA (s,l)
-
-
-
- (** [merge_ac comp l1 l2] merges two lists of terms with coefficients
- into one. Terms that are equal modulo the comparison function
- [comp] will see their arities added. *)
- let merge_ac (compare : 'a -> 'a -> int) (l : 'a mset) (l' : 'a mset) : 'a mset =
- let rec aux l l'=
- match l,l' with
- | [], _ -> l'
- | _, [] -> l
- | (t,tar)::q, (t',tar')::q' ->
- begin match compare t t' with
- | 0 -> ( t,tar+tar'):: aux q q'
- | -1 -> (t, tar):: aux q l'
- | _ -> (t', tar'):: aux l q'
- end
- in aux l l'
-
- (** [merge_map f l] uses the combinator [f] to combine the head of the
- list [l] with the merge_maped tail of [l] *)
- let rec merge_map (f : 'a -> 'b list -> 'b list) (l : 'a list) : 'b list =
- match l with
- | [] -> []
- | t::q -> f t (merge_map f q)
-
-
- (** This function has to deal with the arities *)
- let rec merge l l' =
- merge_ac nf_term_compare l l'
-
- let extract_A s t =
- match t with
- | TA (s',l) when s' = s -> l
- | _ -> [t]
-
- let extract_AC s (t,ar) : nf_term mset =
- match t with
- | TAC (s',l) when s' = s -> List.map (fun (x,y) -> (x,y*ar)) l
- | _ -> [t,ar]
-
-
- (** {3 Constructors of {!nf_term}}*)
-
- let mk_TAC s (l : (nf_term *int) list) =
- mk_TAC' s
- (merge_map (fun u l -> merge (extract_AC s u) l) l)
- let mk_TA s l =
- mk_TA' s
- (merge_map (fun u l -> (extract_A s u) @ l) l)
- let mk_TSym s l = TSym (s,l)
- let mk_TVar v = TVar v
-
-
- (** {2 Printing function} *)
- let print_binary_list single unit binary l =
- let rec aux l =
- match l with
- [] -> unit
- | [t] -> single t
- | t::q ->
- let r = aux q in
- Printf.sprintf "%s" (binary (single t) r)
- in
- aux l
-
- let sprint_ac single (l : 'a mset) =
- (print_binary_list
- (fun (x,t) ->
- if t = 1
- then single x
- else Printf.sprintf "%i*%s" t (single x)
- )
- "0"
- (fun x y -> x ^ " , " ^ y)
- l
- )
-
- let print_symbol single s l =
- match l with
- [] -> Printf.sprintf "%i" s
- | _ ->
- Printf.sprintf "%i(%s)"
- s
- (print_binary_list single "" (fun x y -> x ^ "," ^ y) l)
-
-
- let print_ac_list single s l =
- Printf.sprintf "[%i:AC]{%s}"
- s
- (print_binary_list
- single
- "0"
- (fun x y -> x ^ " , " ^ y)
- l
- )
-
-
- let print_a single s l =
- Printf.sprintf "[%i:A]{%s}"
- s
- (print_binary_list single "1" (fun x y -> x ^ " , " ^ y) l)
-
- let rec sprint_nf_term = function
- | TSym (s,l) -> print_symbol sprint_nf_term s l
- | TAC (s,l) ->
- Printf.sprintf "[%i:AC]{%s}" s
- (sprint_ac
- sprint_nf_term
- l)
- | TA (s,l) -> print_a sprint_nf_term s l
- | TVar v -> Printf.sprintf "x%i" v
-
-
-
- (** {2 Conversion functions} *)
-
- (* rebuilds a tree out of a list *)
- let rec binary_of_list f comb null l =
- let l = List.rev l in
- let rec aux = function
- | [] -> null
- | [t] -> f t
- | t::q -> comb (aux q) (f t)
- in
- aux l
-
- let rec term_of_t : t -> nf_term = function
- | Dot (s,l,r) ->
- let l = term_of_t l in
- let r = term_of_t r in
- mk_TA s [l;r]
- | Plus (s,l,r) ->
- let l = term_of_t l in
- let r = term_of_t r in
- mk_TAC ( s) [l,1;r,1]
- | One x ->
- mk_TA ( x) []
- | Zero x ->
- mk_TAC ( x) []
- | Sym (s,t) ->
- let t = Array.to_list t in
- let t = List.map term_of_t t in
- mk_TSym ( s) t
- | Var i ->
- mk_TVar ( i)
-
- let rec t_of_term : nf_term -> t = function
- | TAC (s,l) ->
- (binary_of_list
- t_of_term
- (fun l r -> Plus ( s,l,r))
- (Zero ( s))
- (linear l)
- )
- | TA (s,l) ->
- (binary_of_list
- t_of_term
- (fun l r -> Dot ( s,l,r))
- (One ( s))
- l
- )
- | TSym (s,l) ->
- let v = Array.of_list l in
- let v = Array.map (t_of_term) v in
- Sym ( s,v)
- | TVar x -> Var x
-
-
- let equal_aac x y =
- nf_equal (term_of_t x) (term_of_t y)
- end
-
-(** Terms environments defined as association lists from variables to
- terms in normal form {! Terms.nf_term} *)
-module Subst : sig
- type t
-
- val find : t -> var -> Terms.nf_term option
- val add : t -> var -> Terms.nf_term -> t
- val empty : t
- val instantiate : t -> Terms.t -> Terms.t
- val sprint : t -> string
- val to_list : t -> (var*Terms.t) list
-end
- =
-struct
- open Terms
-
- (** Terms environments, with nf_terms, to avoid costly conversions
- of {!Terms.nf_terms} to {!Terms.t}, that will be mostly discarded*)
- type t = (var * nf_term) list
-
- let find : t -> var -> nf_term option = fun t x ->
- try Some (List.assoc x t) with | _ -> None
- let add t x v = (x,v) :: t
- let empty = []
-
- let sprint (l : t) =
- match l with
- | [] -> Printf.sprintf "Empty environment\n"
- | _ ->
-
- let s = List.fold_left
- (fun acc (x,y) ->
- Printf.sprintf "%sX%i -> %s\n"
- acc
- x
- (sprint_nf_term y)
- )
- ""
- (List.rev l) in
- Printf.sprintf "%s\n%!" s
-
-
-
- (** [instantiate] is an homomorphism except for the variables*)
- let instantiate (t: t) (x:Terms.t) : Terms.t =
- let rec aux = function
- | One _ as x -> x
- | Zero _ as x -> x
- | Sym (s,t) -> Sym (s,Array.map aux t)
- | Plus (s,l,r) -> Plus (s, aux l, aux r)
- | Dot (s,l,r) -> Dot (s, aux l, aux r)
- | Var i ->
- begin match find t i with
- | None -> Util.error "aac_tactics: instantiate failure"
- | Some x -> t_of_term x
- end
- in aux x
-
- let to_list t = List.map (fun (x,y) -> x,Terms.t_of_term y) t
-end
-
-(******************)
-(* MATCHING UTILS *)
-(******************)
-
-open Terms
-
-(** First, we need to be able to perform non-deterministic choice of
- term splitting to satisfy a pattern. Indeed, we want to show that:
- (x+a*b)*c <= a*b*c
-*)
-let a_nondet_split t : ('a list * 'a list) m =
- let rec aux l l' =
- match l' with
- | [] ->
- return ( l,[])
- | t::q ->
- return ( l,l' )
- >>| aux (l @ [t]) q
- in
- aux [] t
-
-(** Same as the previous [a_nondet_split], but split the list in 3
- parts *)
-let a_nondet_middle t : ('a list * 'a list * 'a list) m =
- a_nondet_split t >>
- (fun left, right ->
- a_nondet_split left >>
- (fun left, middle -> return (left, middle, right))
- )
-
-(** Non deterministic splits of ac lists *)
-let dispatch f n =
- let rec aux k =
- if k = 0 then return (f n 0)
- else return (f (n-k) k) >>| aux (k-1)
- in
- aux (n )
-
-let add_with_arith x ar l =
- if ar = 0 then l else (x,ar) ::l
-
-let ac_nondet_split (l : 'a mset) : ('a mset * 'a mset) m =
- let rec aux = function
- | [] -> return ([],[])
- | (t,tar)::q ->
- aux q
- >>
- (fun (left,right) ->
- dispatch (fun arl arr ->
- add_with_arith t arl left,
- add_with_arith t arr right
- )
- tar
- )
- in
- aux l
-
-(** Try to affect the variable [x] to each left factor of [t]*)
-let var_a_nondet_split ?(strict=false) env current x t =
- a_nondet_split t
- >>
- (fun (l,r) ->
- if strict && l=[] then fail() else
- return ((Subst.add env x (mk_TA current l)), r)
- )
-
-(** Try to affect variable [x] to _each_ subset of t. *)
-let var_ac_nondet_split ?(strict=false) (s : symbol) env (x : var) (t : nf_term mset) : (Subst.t * (nf_term mset)) m =
- ac_nondet_split t
- >>
- (fun (subset,compl) ->
- if strict && subset=[] then fail() else
- return ((Subst.add env x (mk_TAC s subset)), compl)
- )
-
-(** See the term t as a given AC symbol. Unwrap the first constructor
- if necessary *)
-let get_AC (s : symbol) (t : nf_term) : (nf_term *int) list =
- match t with
- | TAC (s',l) when s' = s -> l
- | _ -> [t,1]
-
-(** See the term t as a given A symbol. Unwrap the first constructor
- if necessary *)
-let get_A (s : symbol) (t : nf_term) : nf_term list =
- match t with
- | TA (s',l) when s' = s -> l
- | _ -> [t]
-
-(** See the term [t] as an symbol [s]. Fail if it is not such
- symbol. *)
-let get_Sym s t =
- match t with
- | TSym (s',l) when s' = s -> return l
- | _ -> fail ()
-
-(*************)
-(* A Removal *)
-(*************)
-
-(** We remove the left factor v in a term list. This function runs
- linearly with respect to the size of the first pattern symbol *)
-
-let left_factor current (v : nf_term) (t : nf_term list) =
- let rec aux a b =
- match a,b with
- | t::q , t' :: q' when nf_equal t t' -> aux q q'
- | [], q -> return q
- | _, _ -> fail ()
- in
- match v with
- | TA (s,l) when s = current -> aux l t
- | _ ->
- begin match t with
- | [] -> fail ()
- | t::q ->
- if nf_equal v t
- then return q
- else fail ()
- end
-
-
-(**************)
-(* AC Removal *)
-(**************)
-
-(** [fold_acc] gather all elements of a list that satisfies a
- predicate, and combine them with the residual of the list. That
- is, each element of the residual contains exactly one element less
- than the original term.
-
- TODO : This function not as efficient as it could be
-*)
-
-let pick_sym (s : symbol) (t : nf_term mset ) =
- let rec aux front back =
- match back with
- | [] -> fail ()
- | (t,tar)::q ->
- begin match t with
- | TSym (s',v') when s = s' ->
- let back =
- if tar > 1
- then (t,tar -1) ::q
- else q
- in
- return (v' , List.rev_append front back )
- >>| aux ((t,tar)::front) q
- | _ -> aux ((t,tar)::front) q
- end
- in
- aux [] t
-
-
-
-(** We have to check if we are trying to remove a unit from a term*)
-let is_unit_AC s t =
- nf_equal t (mk_TAC s [])
-
-let is_same_AC s t : nf_term mset option=
- match t with
- TAC (s',l) when s = s' -> Some l
- | _ -> None
-
-(** We want to remove the term [v] from the term list [t] under an AC
- symbol *)
-let single_AC_factor (s : symbol) (v : nf_term) v_ar (t : nf_term mset) : (nf_term mset) m =
- let rec aux front back =
- match back with
- | [] -> fail ()
- | (t,tar)::q ->
- begin
- if nf_equal v t
- then
- match () with
- | _ when tar < v_ar -> fail ()
- | _ when tar = v_ar -> return (List.rev_append front q)
- | _ -> return (List.rev_append front ((t,tar-v_ar)::q))
- else
- aux ((t,tar) :: front) q
- end
- in
- if is_unit_AC s v
- then
- return t
- else
- aux [] t
-
-let factor_AC (s : symbol) (v: nf_term) (t : nf_term mset) : ( nf_term mset ) m =
- match is_same_AC s v with
- | None -> single_AC_factor s v 1 t
- | Some l ->
- (* We are trying to remove an AC factor *)
- List.fold_left (fun acc (v,v_ar) ->
- acc >> (single_AC_factor s v v_ar)
- )
- (return t)
- l
-
-
-
-(************)
-(* Matching *)
-(************)
-
-
-
-(** {!matching} is the generic matching judgement. Each time a
- non-deterministic split is made, we have to come back to this one.
-
- {!matchingSym} is used to match two applications that have the
- same (free) head-symbol.
-
- {!matchingAC} is used to match two sums (with the subtlety that
- [x+y] matches [f a] which is a function application or [a*b] which
- is a product).
-
- {!matchingA} is used to match two products (with the subtlety that
- [x*y] matches [f a] which is a function application, or [a+b]
- which is a sum).
-
-
-*)
-let matching ?strict =
- let rec matching env (p : nf_term) (t: nf_term) : Subst.t Search.m=
- match p with
- | TAC (s,l) ->
- let l = linear l in
- matchingAC env s l (get_AC s t)
- | TA (s,l) ->
- matchingA env s l (get_A s t)
- | TSym (s,l) ->
- (get_Sym s t)
- >> (fun t -> matchingSym env l t)
- | TVar x ->
- begin match Subst.find env x with
- | None -> return (Subst.add env x t)
- | Some v -> if nf_equal v t then return env else fail ()
- end
-
- and
- matchingAC (env : Subst.t) (current : symbol) (l : nf_term list) (t : (nf_term *int) list) =
- match l with
- | TSym (s,v):: q ->
- pick_sym s t
- >> (fun (v',t') ->
- matchingSym env v v'
- >> (fun env -> matchingAC env current q t'))
-
- | TAC (s,v)::q when s = current ->
- assert false
- | TVar x:: q ->
- begin match Subst.find env x with
- | None ->
- (var_ac_nondet_split ?strict current env x t)
- >> (fun (env,residual) -> matchingAC env current q residual)
- | Some v ->
- (factor_AC current v t)
- >> (fun residual -> matchingAC env current q residual)
- end
- | h :: q ->(* PAC =/= curent or PA *)
- (ac_nondet_split t)
- >>
- (fun (left,right) ->
- matching env h (mk_TAC current left)
- >>
- (
- fun env ->
- matchingAC env current q right
- )
- )
- | [] -> if t = [] then return env else fail ()
- and
- matchingA (env : Subst.t) (current : symbol) (l : nf_term list) (t : nf_term list) =
- match l with
- | TSym (s,v) :: l ->
- begin match t with
- | TSym (s',v') :: r when s = s' ->
- (matchingSym env v v')
- >> (fun env -> matchingA env current l r)
- | _ -> fail ()
- end
- | TA (s,v) :: l when s = current ->
- assert false
- | TVar x :: l ->
- begin match Subst.find env x with
- | None ->
- var_a_nondet_split ?strict env current x t
- >> (fun (env,residual)-> matchingA env current l residual)
- | Some v ->
- (left_factor current v t)
- >> (fun residual -> matchingA env current l residual)
- end
- | h :: l ->
- a_nondet_split t
- >> (fun (t,r) ->
- matching env h (mk_TA current t)
- >> (fun env -> matchingA env current l r)
- )
- | [] -> if t = [] then return env else fail ()
- and
- matchingSym (env : Subst.t) (l : nf_term list) (t : nf_term list) =
- List.fold_left2
- (fun acc p t -> acc >> (fun env -> matching env p t))
- (return env)
- l
- t
-
- in
- matching
-
-
-
-(***********)
-(* Subterm *)
-(***********)
-
-(** [tri_fold f l acc] folds on the list [l] and give to f the
- beginning of the list in reverse order, the considered element, and
- the last part of the list
-
- as an exemple, on the list [1;2;3;4], we get the trace
- f () [] 1 [2; 3; 4]
- f () [1] 2 [3; 4]
- f () [2;1] 3 [ 4]
- f () [3;2;1] 4 []
-
- it is the duty of the user to reverse the front if needed
-*)
-
-let tri_fold f (l : 'a list) (acc : 'b)= match l with
- [] -> acc
- | _ ->
- let _,_,acc = List.fold_left (fun acc (t : 'a) ->
- let l,r,acc = acc in
- let r = List.tl r in
- l,r,f acc l t r
- ) ([], l,acc) l
- in acc
-
-
-(** [subterm] solves a sub-term pattern matching.
-
- This function is more high-level than {!matcher}, thus takes {!t}
- as arguments rather than terms in normal form {!nf_term}.
-
- We use three mutually recursive functions {!subterm},
- {!subterm_AC}, {!subterm_A} to find the matching subterm, making
- non-deterministic choices to split the term into a context and an
- intersting sub-term. Intuitively, the only case in which we have to
- go in depth is when we are left with a sub-term that is atomic.
-
- Indeed, rewriting [H: b = c |- a+b+a = a+a+c], we do not want to
- find recursively the sub-terms of [a+b] and [b+a], since they will
- overlap with the sub-terms of [a+b+a].
-
- We rebuild the context on the fly, leaving the variables in the
- pattern uninstantiated. We do so in order to allow interaction
- with the user, to choose the env.
-
- Strange patterms like x*y*x can be instanciated by nothing, inside
- a product. Therefore, we need to check that all the term is not
- going into the context (hence the tests on the length of the
- lists). With proper support for interaction with the user, we
- should lift these tests. However, at the moment, they serve as
- heuristics to return "interesting" matchings
-
-*)
-
-let return_non_empty raw_p m =
- if Search.is_empty m
- then
- fail ()
- else
- return (raw_p ,m)
-
-let subterm ?strict (raw_p:t) (raw_t:t): (int* t * Subst.t m) m=
- let p = term_of_t raw_p in
- let t = term_of_t raw_t in
- let rec subterm (t:nf_term) : (t * Subst.t m) m=
- match t with
- | TAC (s,l) ->
- (ac_nondet_split l) >>
- (fun (left,right) ->
- (subterm_AC s left) >>
- (fun (p,m) ->
- let p = if right = [] then p else
- Plus (s,p,t_of_term (mk_TAC s right))
- in
- return (p,m)
- )
-
- )
- | TA (s,l) ->
- (a_nondet_middle l)
- >>
- (fun (left, middle, right) ->
- (subterm_A s middle) >>
- (fun (p,m) ->
- let p =
- if right = [] then p else
- Dot (s,p,t_of_term (mk_TA s right))
- in
- let p =
- if left = [] then p else
- Dot (s,t_of_term (mk_TA s left),p)
- in
- return (p,m)
- )
- )
- | TSym (s, l) ->
- let init = return_non_empty raw_p (matching ?strict Subst.empty p t) in
- tri_fold (fun acc l t r ->
- ((subterm t) >>
- (fun (p,m) ->
- let l = List.map t_of_term l in
- let r = List.map t_of_term r in
- let p = Sym (s, Array.of_list (List.rev_append l (p::r))) in
- return (p,m)
- )) >>| acc
- ) l init
- | _ -> assert false
- and subterm_AC s tl =
- match tl with
- [x,1] -> subterm x
- | _ ->
- return_non_empty raw_p (matching ?strict Subst.empty p (mk_TAC s tl))
- and subterm_A s tl =
- match tl with
- [x] -> subterm x
- | _ ->
- return_non_empty raw_p (matching ?strict Subst.empty p (mk_TA s tl))
- in
- (subterm t >> fun (p,m) -> return (Terms.size p,p,m))
-
-(* The functions we export, handlers for the previous ones. Some debug
- information also *)
-let subterm ?strict raw t =
- let sols = time (subterm ?strict raw) t "%fs spent in subterm (including matching)\n" in
- if debug then Printf.printf "%i possible solution(s)\n"
- (Search.fold (fun (_,_,envm) acc -> count envm + acc) sols 0);
- sols
-
-
-let matcher ?strict p t =
- let sols = time
- (fun (p,t) ->
- let p = (Terms.term_of_t p) in
- let t = (Terms.term_of_t t) in
- matching ?strict Subst.empty p t) (p,t)
- "%fs spent in the matcher\n"
- in
- if debug then Printf.printf "%i solutions\n" (count sols);
- sols
-
-(* A very basic way to interact with the envs, to choose a possible
- solution *)
-open Pp
-let pp_env pt : Subst.t -> Pp.std_ppcmds = fun env ->
- List.fold_left (fun acc (v,t) -> str (Printf.sprintf "x%i: " v) ++ pt t ++ str "; " ++ acc) (str "") (Subst.to_list env)
-
-let pp_envm pt : Subst.t Search.m -> Pp.std_ppcmds = fun m ->
- let _,s = Search.fold
- (fun env (n,acc) ->
- n+1, h 0 (str (Printf.sprintf "%i:\t[" n) ++pp_env pt env ++ str "]") ++ fnl () :: acc
- ) m (0,[]) in
- List.fold_left (fun acc s -> s ++ acc) (str "") (s)
-
-let pp_all pt : (int * Terms.t * Subst.t Search.m) Search.m -> Pp.std_ppcmds = fun m ->
- let _,s = Search.fold
- (fun (size,context,envm) (n,acc) ->
- let s = str (Printf.sprintf "subterm %i\t" n) in
- let s = s ++ (str "(context ") ++ pt context ++ (str ")\n") in
- let s = s ++ str (Printf.sprintf "\t%i possible(s) substitutions" (Search.count envm) ) ++ fnl () in
- let s = s ++ pp_envm pt envm in
- n+1, s::acc
- ) m (0,[]) in
- List.fold_left (fun acc s -> s ++ str "\n" ++ acc) (str "") (s)
-
diff --git a/theory.ml b/theory.ml
deleted file mode 100644
index 45a76f1..0000000
--- a/theory.ml
+++ /dev/null
@@ -1,708 +0,0 @@
-(***************************************************************************)
-(* This is part of aac_tactics, it is distributed under the terms of the *)
-(* GNU Lesser General Public License version 3 *)
-(* (see file LICENSE for more details) *)
-(* *)
-(* Copyright 2009-2010: Thomas Braibant, Damien Pous. *)
-(***************************************************************************)
-
-(** Constr from the theory of this particular development
-
- The inner-working of this module is highly correlated with the
- particular structure of {b AAC.v}, therefore, it should
- be of little interest to most readers.
-*)
-
-
-open Term
-
-let ac_theory_path = ["AAC"]
-
-module Sigma = struct
- let sigma = lazy (Coq.init_constant ac_theory_path "sigma")
- let sigma_empty = lazy (Coq.init_constant ac_theory_path "sigma_empty")
- let sigma_add = lazy (Coq.init_constant ac_theory_path "sigma_add")
- let sigma_get = lazy (Coq.init_constant ac_theory_path "sigma_get")
-
- let add ty n x map =
- mkApp (Lazy.force sigma_add,[|ty; n; x ; map|])
- let empty ty =
- mkApp (Lazy.force sigma_empty,[|ty |])
- let typ ty =
- mkApp (Lazy.force sigma, [|ty|])
-
- let to_fun ty null map =
- mkApp (Lazy.force sigma_get, [|ty;null;map|])
-
- let of_list ty null l =
- let map =
- List.fold_left
- (fun acc (i,t) -> assert (i > 0); add ty (Coq.Pos.of_int i) t acc)
- (empty ty)
- l
- in to_fun ty null map
-end
-
-
-module Sym = struct
- type pack = {ar: Term.constr; value: Term.constr ; morph: Term.constr}
-
- let path = ac_theory_path @ ["Sym"]
- let typ = lazy (Coq.init_constant path "pack")
- let mkPack = lazy (Coq.init_constant path "mkPack")
- let typeof = lazy (Coq.init_constant path "type_of")
- let relof = lazy (Coq.init_constant path "rel_of")
- let value = lazy (Coq.init_constant path "value")
- let null = lazy (Coq.init_constant path "null")
- let mk_typeof : Coq.reltype -> int -> constr = fun (x,r) n ->
- mkApp (Lazy.force typeof, [| x ; Coq.Nat.of_int n |])
- let mk_relof : Coq.reltype -> int -> constr = fun (x,r) n ->
- mkApp (Lazy.force relof, [| x;r ; Coq.Nat.of_int n |])
- let mk_pack ((x,r): Coq.reltype) s =
- mkApp (Lazy.force mkPack, [|x;r; s.ar;s.value;s.morph|])
- let null eqt witness =
- let (x,r),e = eqt in
- mkApp (Lazy.force null, [| x;r;e;witness |])
-
-end
-
-module Sum = struct
- type pack = {plus: Term.constr; zero: Term.constr ; opac: Term.constr}
-
- let path = ac_theory_path @ ["Sum"]
- let typ = lazy (Coq.init_constant path "pack")
- let mkPack = lazy (Coq.init_constant path "mkPack")
- let cstr_zero = lazy (Coq.init_constant path "zero")
- let cstr_plus = lazy (Coq.init_constant path "plus")
-
- let mk_pack: Coq.reltype -> pack -> Term.constr = fun (x,r) s ->
- mkApp (Lazy.force mkPack , [| x ; r; s.plus; s.zero; s.opac|])
-
- let zero: Coq.reltype -> constr -> constr = fun (x,r) pack ->
- mkApp (Lazy.force cstr_zero, [| x;r;pack|])
- let plus: Coq.reltype -> constr -> constr = fun (x,r) pack ->
- mkApp (Lazy.force cstr_plus, [| x;r;pack|])
-
-
-end
-
-module Prd = struct
- type pack = {dot: Term.constr; one: Term.constr ; opa: Term.constr}
-
- let path = ac_theory_path @ ["Prd"]
- let typ = lazy (Coq.init_constant path "pack")
- let mkPack = lazy (Coq.init_constant path "mkPack")
- let cstr_one = lazy (Coq.init_constant path "one")
- let cstr_dot = lazy (Coq.init_constant path "dot")
-
- let mk_pack: Coq.reltype -> pack -> Term.constr = fun (x,r) s ->
- mkApp (Lazy.force mkPack , [| x ; r; s.dot; s.one; s.opa|])
- let one: Coq.reltype -> constr -> constr = fun (x,r) pack ->
- mkApp (Lazy.force cstr_one, [| x;r;pack|])
- let dot: Coq.reltype -> constr -> constr = fun (x,r) pack ->
- mkApp (Lazy.force cstr_dot, [| x;r;pack|])
-
-
-end
-
-
-let op_ac = lazy (Coq.init_constant ac_theory_path "Op_AC")
-let op_a = lazy (Coq.init_constant ac_theory_path "Op_A")
-
-(** The constants from the inductive type *)
-let _Tty = lazy (Coq.init_constant ac_theory_path "T")
-let vTty = lazy (Coq.init_constant ac_theory_path "vT")
-
-let rsum = lazy (Coq.init_constant ac_theory_path "sum")
-let rprd = lazy (Coq.init_constant ac_theory_path "prd")
-let rsym = lazy (Coq.init_constant ac_theory_path "sym")
-let vnil = lazy (Coq.init_constant ac_theory_path "vnil")
-let vcons = lazy (Coq.init_constant ac_theory_path "vcons")
-let eval = lazy (Coq.init_constant ac_theory_path "eval")
-let decide_thm = lazy (Coq.init_constant ac_theory_path "decide")
-
-(** Rebuild an equality *)
-let mk_equal = fun (x,r) left right ->
- mkApp (r, [| left; right|])
-
-
-let anomaly msg =
- Util.anomaly ("aac_tactics: " ^ msg)
-
-let user_error msg =
- Util.error ("aac_tactics: " ^ msg)
-
-
-module Trans = struct
-
-
- (** {1 From Coq to Abstract Syntax Trees (AST)}
-
- This module provides facilities to interpret a Coq term with
- arbitrary operators as an abstract syntax tree. Considering an
- application, we try to infer instances of our classes. We
- consider that raw morphisms ([Sym]) are coarser than [A]
- operators, which in turn are coarser than [AC] operators. We
- need to treat the units first. Indeed, considering [S] as a
- morphism is problematic because [S O] is a unit.
-
- During this reification, we gather some informations that will
- be used to rebuild Coq terms from AST ( type {!envs}) *)
-
- type pack =
- | AC of Sum.pack (* the opac pack *)
- | A of Prd.pack (* the opa pack *)
- | Sym of Sym.pack (* the sym pack *)
-
- type head =
- | Fun
- | Plus
- | Dot
- | One
- | Zero
-
-
- (* discr is a map from {!Term.constr} to
- (head * pack).
-
- [None] means that it is not possible to instantiate this partial
- application to an interesting class.
-
- [Some x] means that we found something in the past. This means in
- particular that a single constr cannot be two things at the same
- time.
-
- However, for the transitivity process, we do not need to remember
- if the constr was the unit or the binary operation. We will use
- the projections from the module's record.
-
- The field [bloom] allows to give a unique number to each class we
- found. *)
- type envs =
- {
- discr : (constr , (head * pack) option) Hashtbl.t ;
-
- bloom : (pack, int ) Hashtbl.t;
- bloom_back : (int, pack ) Hashtbl.t;
- bloom_next : int ref;
- }
-
- let empty_envs () =
- {
- discr = Hashtbl.create 17;
-
- bloom = Hashtbl.create 17;
- bloom_back = Hashtbl.create 17;
- bloom_next = ref 1;
- }
-
- let add_bloom envs pack =
- if Hashtbl.mem envs.bloom pack
- then ()
- else
- let x = ! (envs.bloom_next) in
- Hashtbl.add envs.bloom pack x;
- Hashtbl.add envs.bloom_back x pack;
- incr (envs.bloom_next)
-
- let find_bloom envs pack =
- try Hashtbl.find envs.bloom pack
- with Not_found -> assert false
-
-
- (** try to infer the kind of operator we are dealing with *)
- let is_morphism rlt papp ar goal =
- let typeof = Sym.mk_typeof rlt ar in
- let relof = Sym.mk_relof rlt ar in
- let env = Tacmach.pf_env goal in
- let evar_map = Tacmach.project goal in
- let m = Coq.Classes.mk_morphism typeof relof papp in
- try
- let _ = Tacmach.pf_check_type goal papp typeof in
- let evar_map,m = Typeclasses.resolve_one_typeclass env evar_map m in
- let pack = {Sym.ar = (Coq.Nat.of_int ar); Sym.value= papp; Sym.morph= m} in
- Some (Coq.goal_update goal evar_map, Fun , Sym pack)
- with
- | Not_found -> None
- | _ -> None
-
- (** gives back the class *)
- let is_op gl (rlt: Coq.reltype) pack op null =
- let (x,r) = rlt in
- let ty = mkApp (pack ,[|x;r; op; null |]) in
- let env = Tacmach.pf_env gl in
- let evar_map = Tacmach.project gl in
- try
- let em,t = Typeclasses.resolve_one_typeclass env evar_map ty in
- let op = Evarutil.nf_evar em op in
- let null = Evarutil.nf_evar em null in
- Some (Coq.goal_update gl em,t,op,null)
- with Not_found -> None
-
-
- let is_plus (rlt: Coq.reltype) plus goal=
- let (zero,goal) = Coq.evar_unit goal rlt in
- match is_op goal rlt (Lazy.force op_ac) plus zero
- with
- | None -> None
- | Some (gl,s,plus,zero) ->
- let pack = {Sum.plus=plus;Sum.zero=zero;Sum.opac=s} in
- Some (gl, Plus , AC pack)
-
- let is_dot rlt dot goal=
- let (one,goal) = Coq.evar_unit goal rlt in
- match is_op goal rlt (Lazy.force op_a) dot one
- with
- | None -> None
- | Some (goal,s,dot,one) ->
- let pack = {Prd.dot=dot;Prd.one=one;Prd.opa=s} in
- Some (goal, Dot, A pack)
-
- let is_zero rlt zero goal =
- let (plus,goal) = Coq.evar_binary goal rlt in
- match is_op goal rlt (Lazy.force op_ac) plus zero
- with
- | None -> None
- | Some (goal,s,plus,zero) ->
- let pack = {Sum.plus=plus;Sum.zero=zero;Sum.opac=s} in
- Some (goal, Zero , AC pack)
-
- let is_one rlt one goal =
- let (dot,goal) = Coq.evar_binary goal rlt in
- match is_op goal rlt (Lazy.force op_a) dot one
- with
- | None -> None
- | Some (goal,s,dot,one)->
- let pack = {Prd.dot=dot;Prd.one=one;Prd.opa=s} in
- Some (goal, One, A pack)
-
-
- let (>>) x y a =
- match x a with | None -> y a | x -> x
-
- let is_unit rlt papp =
- (is_zero rlt papp )>> (is_one rlt papp)
-
- let what_is_it rlt papp ar goal :
- (Coq.goal_sigma * head * pack ) option =
- match ar with
- | 2 ->
- begin
- (
- (is_plus rlt papp)>>
- (is_dot rlt papp)>>
- (is_morphism rlt papp ar)
- ) goal
- end
- | _ ->
- is_morphism rlt papp ar goal
-
-
-
-
- (** [discriminates goal envs rlt t ca] infer :
-
- - the nature of the application [t ca] (is the head symbol a
- Plus, a Dot, so on),
-
- - its {! pack } (is it an AC operator, or an A operator, or a
- Symbol ?),
-
- - the relevant partial application,
-
- - the vector of the remaining arguments
-
- We use an expansion to handle the special case of units, before
- going back to the general discrimination procedure *)
- let discriminate goal envs (rlt : Coq.reltype) t ca : Coq.goal_sigma * head * pack * constr * constr array =
- let nb_params = Array.length ca in
- let f cont x p_app ar =
- assert (0 <= ar && ar <= nb_params);
- match x with
- | Some (goal, discr, pack) ->
- Hashtbl.add envs.discr p_app (Some (discr, pack));
- add_bloom envs pack;
- (goal, discr, pack, p_app, Array.sub ca (nb_params-ar) ar)
- | None ->
- begin
- (* to memoise the failures *)
- Hashtbl.add envs.discr p_app None;
- (* will terminate, since [const] is
- capped, and it is easy to find an
- instance of a constant *)
- cont goal (ar-1)
- end
- in
- let rec aux goal ar :Coq.goal_sigma * head * pack * constr * constr array =
- assert (0 <= ar && ar <= nb_params);
- let p_app = mkApp (t, Array.sub ca 0 (nb_params - ar)) in
- begin
- try
- begin match Hashtbl.find envs.discr p_app with
- | None -> aux goal (ar-1)
- | Some (discr,sym) ->
- goal,
- discr,
- sym,
- p_app,
- Array.sub ca (nb_params -ar ) ar
- end
- with
- Not_found -> (* Could not find this constr *)
- f aux (what_is_it rlt p_app ar goal) p_app ar
- end
- in
- match is_unit rlt (mkApp (t,ca)) goal with
- | None -> aux goal (nb_params)
- | x -> f (fun _ -> assert false) x (mkApp (t,ca)) 0
-
- let discriminate goal envs rlt x =
- try
- match Coq.decomp_term x with
- | Var _ ->
- discriminate goal envs rlt x [| |]
- | App (t,ca) ->
- discriminate goal envs rlt t ca
- | Construct c ->
- discriminate goal envs rlt x [| |]
- | _ ->
- assert false
- with
- (* TODO: is it the only source of invalid use that fall into
- this catch_all ? *)
- e -> user_error "cannot handle this kind of hypotheses (higher order function symbols, variables occuring under a non-morphism, etc)."
-
-
- (** [t_of_constr goal rlt envs cstr] builds the abstract syntax
- tree of the term [cstr]. Doing so, it modifies the environment
- of known stuff [envs], and eventually creates fresh
- evars. Therefore, we give back the goal updated accordingly *)
- let t_of_constr goal (rlt: Coq.reltype ) envs : constr -> Matcher.Terms.t *Coq.goal_sigma =
- let r_evar_map = ref (goal) in
- let rec aux x =
- match Coq.decomp_term x with
- | Rel i -> Matcher.Terms.Var i
- | _ ->
- let goal, sym, pack , p_app, ca = discriminate (!r_evar_map) envs rlt x in
- r_evar_map := goal;
- let k = find_bloom envs pack
- in
- match sym with
- | Plus ->
- assert (Array.length ca = 2);
- Matcher.Terms.Plus ( k, aux ca.(0), aux ca.(1))
- | Zero ->
- assert (Array.length ca = 0);
- Matcher.Terms.Zero ( k)
- | Dot ->
- assert (Array.length ca = 2);
- Matcher.Terms.Dot ( k, aux ca.(0), aux ca.(1))
- | One ->
- assert (Array.length ca = 0);
- Matcher.Terms.One ( k)
- | Fun ->
- Matcher.Terms.Sym ( k, Array.map aux ca)
- in
- (
- fun x -> let r = aux x in r, !r_evar_map
- )
-
- (** {1 From AST back to Coq }
-
- The next functions allow one to map OCaml abstract syntax trees
- to Coq terms *)
-
- (** {2 Building raw, natural, terms} *)
-
- (** [raw_constr_of_t] rebuilds a term in the raw representation *)
- let raw_constr_of_t envs (rlt : Coq.reltype) (t : Matcher.Terms.t) : Term.constr =
- let add_set,get =
- let r = ref [] in
- let rec add x = function
- [ ] -> [x]
- | t::q when t = x -> t::q
- | t::q -> t:: (add x q)
- in
- (fun x -> r := add x !r),(fun () -> !r)
- in
- let rec aux t =
- match t with
- | Matcher.Terms.Plus (s,l,r) ->
- begin match Hashtbl.find envs.bloom_back s with
- | AC s ->
- mkApp (s.Sum.plus , [|(aux l);(aux r)|])
- | _ -> assert false
- end
- | Matcher.Terms.Dot (s,l,r) ->
- begin match Hashtbl.find envs.bloom_back s with
- | A s ->
- mkApp (s.Prd.dot, [|(aux l);(aux r)|])
- | _ -> assert false
- end
- | Matcher.Terms.Sym (s,t) ->
- begin match Hashtbl.find envs.bloom_back s with
- | Sym s ->
- mkApp (s.Sym.value, Array.map aux t)
- | _ -> assert false
- end
- | Matcher.Terms.One x ->
- begin match Hashtbl.find envs.bloom_back x with
- | A s ->
- s.Prd.one
- | _ -> assert false
- end
- | Matcher.Terms.Zero x ->
- begin match Hashtbl.find envs.bloom_back x with
- | AC s ->
- s.Sum.zero
- | _ -> assert false
- end
- | Matcher.Terms.Var i -> add_set i;
- let i = (Names.id_of_string (Printf.sprintf "x%i" i)) in
- mkVar (i)
- in
- let x = fst rlt in
- let rec cap c = function [] -> c
- | t::q -> let i = (Names.id_of_string (Printf.sprintf "x%i" t)) in
- cap (mkNamedProd i x c) q
- in
- cap ((aux t)) (get ())
- (* aux t is possible for printing only, but I wonder what
- would be the status of such a term*)
-
-
- (** {2 Building reified terms} *)
-
- (* Some informations to be able to build the maps *)
- type reif_params =
- {
- a_null : constr;
- ac_null : constr;
- sym_null : constr;
- a_ty : constr;
- ac_ty : constr;
- sym_ty : constr;
- }
-
- type sigmas = {
- env_sym : Term.constr;
- env_sum : Term.constr;
- env_prd : Term.constr;
- }
- (** A record containing the environments that will be needed by the
- decision procedure, as a Coq constr. Contains also the functions
- from the symbols (as ints) to indexes (as constr) *)
- type sigma_maps = {
- sym_to_pos : int -> Term.constr;
- sum_to_pos : int -> Term.constr;
- prd_to_pos : int -> Term.constr;
- }
-
- (* convert the [envs] into the [sigma_maps] *)
- let envs_to_list rlt envs =
- let add x y (n,l) = (n+1, (x, ( n, y))::l) in
- let nil = 1,[]in
- Hashtbl.fold
- (fun int pack (sym,sum,prd) ->
- match pack with
- |AC s ->
- let s = Sum.mk_pack rlt s in
- sym, add (int) s sum, prd
- |A s ->
- let s = Prd.mk_pack rlt s in
- sym, sum, add (int) s prd
- | Sym s ->
- let s = Sym.mk_pack rlt s in
- add (int) s sym, sum, prd
- )
- envs.bloom_back
- (nil,nil,nil)
-
-
-
- (** infers some stuff that will be required when we will build
- environments (our environments need a default case, so we need
- an Op_AC, an Op_A, and a symbol) *)
-
- (* Note : this function can fail if the user is using the wrong
- relation, like proving a = b, while the classes are defined with
- another relation (==)
- *)
- let build_reif_params goal eqt =
- let rlt = fst eqt in
- let plus,goal = Coq.evar_binary goal rlt in
- let dot ,goal = Coq.evar_binary goal rlt in
- let one ,goal = Coq.evar_unit goal rlt in
- let zero,goal = Coq.evar_unit goal rlt in
-
- let (x,r) = rlt in
- let ty_ac = ( mkApp (Lazy.force op_ac, [| x;r;plus;zero|])) in
- let ty_a = ( mkApp (Lazy.force op_a, [| x;r;dot;one|])) in
- let a ,goal=
- try Coq.resolve_one_typeclass goal ty_a
- with Not_found -> user_error ("Unable to infer a default A operator.")
- in
- let ac ,goal=
- try Coq.resolve_one_typeclass goal ty_ac
- with Not_found -> user_error ("Unable to infer a default AC operator.")
- in
- let plus = Coq.nf_evar goal plus in
- let dot = Coq.nf_evar goal dot in
- let one = Coq.nf_evar goal one in
- let zero = Coq.nf_evar goal zero in
-
- let record =
- {a_null = Prd.mk_pack rlt {Prd.dot=dot;Prd.one=one;Prd.opa=a};
- ac_null= Sum.mk_pack rlt {Sum.plus=plus;Sum.zero=zero;Sum.opac=ac};
- sym_null = Sym.null eqt zero;
- a_ty = (mkApp (Lazy.force Prd.typ, [| x;r; |]));
- ac_ty = ( mkApp (Lazy.force Sum.typ, [| x;r|]));
- sym_ty = mkApp (Lazy.force Sym.typ, [|x;r|])
- }
- in
- goal, record
-
- (** [build_sigma_maps] given a envs and some reif_params, we are
- able to build the sigmas *)
- let build_sigma_maps goal eqt envs : sigmas * sigma_maps * Proof_type.goal Evd.sigma =
- let rlt = fst eqt in
- let goal,rp = build_reif_params goal eqt in
- let ((_,sym), (_,sum),(_,prd)) = envs_to_list rlt envs in
- let f = List.map (fun (x,(y,z)) -> (x,Coq.Pos.of_int y)) in
- let g = List.map (fun (x,(y,z)) -> (y,z)) in
- let sigmas =
- {env_sym = Sigma.of_list rp.sym_ty rp.sym_null (g sym);
- env_sum = Sigma.of_list rp.ac_ty rp.ac_null (g sum);
- env_prd = Sigma.of_list rp.a_ty rp.a_null (g prd);
- } in
- let sigma_maps = {
- sym_to_pos = (let sym = f sym in fun x -> (List.assoc x sym));
- sum_to_pos = (let sum = f sum in fun x -> (List.assoc x sum));
- prd_to_pos = (let prd = f prd in fun x -> (List.assoc x prd));
- }
- in
- sigmas, sigma_maps , goal
-
-
-
- (** builders for the reification *)
- type reif_builders =
- {
- rsum: constr -> constr -> constr -> constr;
- rprd: constr -> constr -> constr -> constr;
- rzero: constr -> constr;
- rone: constr -> constr;
- rsym: constr -> constr array -> constr
- }
-
- (* donne moi une tactique, je rajoute ma part. Potentiellement, il
- est possible d'utiliser la notation 'do' a la Haskell:
- http://www.cas.mcmaster.ca/~carette/pa_monad/ *)
- let (>>) : 'a * Proof_type.tactic -> ('a -> 'b * Proof_type.tactic) -> 'b * Proof_type.tactic =
- fun (x,t) f ->
- let b,t' = f x in
- b, Tacticals.tclTHEN t t'
-
- let return x = x, Tacticals.tclIDTAC
-
- let mk_vect vnil vcons v =
- let ar = Array.length v in
- let rec aux = function
- | 0 -> vnil
- | n -> let n = n-1 in
- mkApp( vcons,
- [|
- (Coq.Nat.of_int n);
- v.(ar - 1 - n);
- (aux (n))
- |]
- )
- in aux ar
-
- (* TODO: use a do notation *)
- let mk_reif_builders (rlt: Coq.reltype) (env_sym:constr) :reif_builders*Proof_type.tactic =
- let x,r = rlt in
- let tty = mkApp (Lazy.force _Tty, [| x ; r ; env_sym|]) in
- (Coq.mk_letin' "mk_rsum" (mkApp (Lazy.force rsum, [| x; r; env_sym|]))) >>
- (fun rsum -> (Coq.mk_letin' "mk_rprd" (mkApp (Lazy.force rprd, [| x; r; env_sym|])))
- >> fun rprd -> ( Coq.mk_letin' "mk_rsym" (mkApp (Lazy.force rsym, [| x;r; env_sym|])))
- >> fun rsym -> (Coq.mk_letin' "mk_vnil" (mkApp (Lazy.force vnil, [| x;r; env_sym|])))
- >> fun vnil -> Coq.mk_letin' "mk_vcons" (mkApp (Lazy.force vcons, [| x;r; env_sym|]))
- >> fun vcons ->
- return {
- rsum =
- begin fun idx l r ->
- mkApp (rsum, [| idx ; Coq.mk_mset tty [l,1 ; r,1]|])
- end;
- rzero =
- begin fun idx ->
- mkApp (rsum, [| idx ; Coq.mk_mset tty []|])
- end;
- rprd =
- begin fun idx l r ->
- let lst = Coq.List.of_list tty [l;r] in
- mkApp (rprd, [| idx; lst|])
- end;
- rone =
- begin fun idx ->
- let lst = Coq.List.of_list tty [] in
- mkApp (rprd, [| idx; lst|]) end;
- rsym =
- begin fun idx v ->
- let vect = mk_vect vnil vcons v in
- mkApp (rsym, [| idx; vect|])
- end;
- }
- )
-
-
-
- type reifier = sigma_maps * reif_builders
-
- let mk_reifier eqt envs goal : sigmas * reifier * Proof_type.tactic =
- let s, sm, goal = build_sigma_maps goal eqt envs in
- let env_sym, env_sym_letin = Coq.mk_letin "env_sym_name" (s.env_sym) in
- let env_prd, env_prd_letin = Coq.mk_letin "env_prd_name" (s.env_prd) in
- let env_sum, env_sum_letin = Coq.mk_letin "env_sum_name" (s.env_sum) in
- let s = {env_sym = env_sym; env_prd = env_prd; env_sum = env_sum} in
- let rb , tac = mk_reif_builders (fst eqt) (s.env_sym) in
- let evar_map = Tacmach.project goal in
- let tac = Tacticals.tclTHENLIST (
- [ Refiner.tclEVARS evar_map;
- env_prd_letin ;
- env_sum_letin ;
- env_sym_letin ;
- tac
- ]) in
- s, (sm, rb), tac
-
-
- (** [reif_constr_of_t reifier t] rebuilds the term [t] in the
- reified form. We use the [reifier] to minimise the size of the
- terms (we make as much lets as possible)*)
- let reif_constr_of_t (sm,rb) (t:Matcher.Terms.t) : constr =
- let rec aux = function
- | Matcher.Terms.Plus (s,l,r) ->
- let idx = sm.sum_to_pos s in
- rb.rsum idx (aux l) (aux r)
- | Matcher.Terms.Dot (s,l,r) ->
- let idx = sm.prd_to_pos s in
- rb.rprd idx (aux l) (aux r)
- | Matcher.Terms.Sym (s,t) ->
- let idx = sm.sym_to_pos s in
- rb.rsym idx (Array.map aux t)
- | Matcher.Terms.One s ->
- let idx = sm.prd_to_pos s in
- rb.rone idx
- | Matcher.Terms.Zero s ->
- let idx = sm.sum_to_pos s in
- rb.rzero idx
- | Matcher.Terms.Var i ->
- anomaly "call to reif_constr_of_t on a term with variables."
- in aux t
-
-
-end
-
-
-