From c0a3544d6351e19c695951796bcee838671d1098 Mon Sep 17 00:00:00 2001 From: letouzey Date: Thu, 5 May 2011 15:12:15 +0000 Subject: Modularization of BinPos + fixes in Stdlib BinPos now contain a sub-module Pos, in which are placed functions like add (ex-Pplus), mul (ex-Pmult), ... and properties like add_comm, add_assoc, ... In addition to the name changes, the organisation is changed quite a lot, to try to take advantage more of the orders < and <= instead of speaking only of the comparison function. The main source of incompatibilities in scripts concerns this compare: Pos.compare is now a binary operation, expressed in terms of the ex-Pcompare which is ternary (expecting an initial comparision as 3rd arg), this ternary version being called now Pos.compare_cont. As for everything else, compatibility notations (only parsing) are provided. But notations "_ ?= _" on positive will have to be edited, since they now point to Pos.compare. We also make the sub-module Pos to be directly an OrderedType, and include results about min and max. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@14098 85f007b7-540e-0410-9357-904b9bb8a0f7 --- plugins/rtauto/Bintree.v | 162 ++++++++--------------------------------------- plugins/rtauto/Rtauto.v | 37 ++++++----- 2 files changed, 43 insertions(+), 156 deletions(-) (limited to 'plugins/rtauto') diff --git a/plugins/rtauto/Bintree.v b/plugins/rtauto/Bintree.v index 6c250ef4f..e3e9078e3 100644 --- a/plugins/rtauto/Bintree.v +++ b/plugins/rtauto/Bintree.v @@ -12,139 +12,20 @@ Require Export BinPos. Open Scope positive_scope. Ltac clean := try (simpl; congruence). -Ltac caseq t := generalize (refl_equal t); pattern t at -1; case t. - -Functional Scheme Pcompare_ind := Induction for Pcompare Sort Prop. - -Lemma Gt_Eq_Gt : forall p q cmp, - (p ?= q) Eq = Gt -> (p ?= q) cmp = Gt. -apply (Pcompare_ind (fun p q cmp _ => (p ?= q) Eq = Gt -> (p ?= q) cmp = Gt)); -simpl;auto;congruence. -Qed. - -Lemma Gt_Lt_Gt : forall p q cmp, - (p ?= q) Lt = Gt -> (p ?= q) cmp = Gt. -apply (Pcompare_ind (fun p q cmp _ => (p ?= q) Lt = Gt -> (p ?= q) cmp = Gt)); -simpl;auto;congruence. -Qed. - -Lemma Gt_Psucc_Eq: forall p q, - (p ?= Psucc q) Gt = Gt -> (p ?= q) Eq = Gt. -intros p q;generalize p;clear p;induction q;destruct p;simpl;auto;try congruence. -intro;apply Gt_Eq_Gt;auto. -apply Gt_Lt_Gt. -Qed. - -Lemma Eq_Psucc_Gt: forall p q, - (p ?= Psucc q) Eq = Eq -> (p ?= q) Eq = Gt. -intros p q;generalize p;clear p;induction q;destruct p;simpl;auto;try congruence. -intro H;elim (Pcompare_not_Eq p (Psucc q));tauto. -intro H;apply Gt_Eq_Gt;auto. -intro H;rewrite Pcompare_Eq_eq with p q;auto. -generalize q;clear q IHq p H;induction q;simpl;auto. -intro H;elim (Pcompare_not_Eq p q);tauto. -Qed. - -Lemma Gt_Psucc_Gt : forall n p cmp cmp0, - (n?=p) cmp = Gt -> (Psucc n?=p) cmp0 = Gt. -induction n;intros [ | p | p];simpl;try congruence. -intros; apply IHn with cmp;trivial. -intros; apply IHn with Gt;trivial. -intros;apply Gt_Lt_Gt;trivial. -intros [ | | ] _ H. -apply Gt_Eq_Gt;trivial. -apply Gt_Lt_Gt;trivial. -trivial. -Qed. Lemma Gt_Psucc: forall p q, - (p ?= Psucc q) Eq = Gt -> (p ?= q) Eq = Gt. -intros p q;generalize p;clear p;induction q;destruct p;simpl;auto;try congruence. -apply Gt_Psucc_Eq. -intro;apply Gt_Eq_Gt;apply IHq;auto. -apply Gt_Eq_Gt. -apply Gt_Lt_Gt. + (p ?= Psucc q) = Gt -> (p ?= q) = Gt. +Proof. +intros. rewrite <- Pos.compare_succ_succ. +now apply Pos.lt_gt, Pos.lt_lt_succ, Pos.gt_lt. Qed. Lemma Psucc_Gt : forall p, - (Psucc p ?= p) Eq = Gt. -induction p;simpl. -apply Gt_Eq_Gt;auto. -generalize p;clear p IHp. -induction p;simpl;auto. -reflexivity. -Qed. - -Fixpoint pos_eq (m n:positive) {struct m} :bool := -match m, n with - xI mm, xI nn => pos_eq mm nn -| xO mm, xO nn => pos_eq mm nn -| xH, xH => true -| _, _ => false -end. - -Theorem pos_eq_refl : forall m n, pos_eq m n = true -> m = n. -induction m;simpl;intro n;destruct n;congruence || -(intro e;apply f_equal;auto). -Defined. - -Theorem refl_pos_eq : forall m, pos_eq m m = true. -induction m;simpl;auto. + (Psucc p ?= p) = Gt. +Proof. +intros. apply Pos.lt_gt, Pos.lt_succ_diag_r. Qed. -Definition pos_eq_dec : forall (m n:positive), {m=n}+{m<>n} . -fix 1;intros [mm|mm|] [nn|nn|];try (right;congruence). -case (pos_eq_dec mm nn). -intro e;left;apply (f_equal xI e). -intro ne;right;congruence. -case (pos_eq_dec mm nn). -intro e;left;apply (f_equal xO e). -intro ne;right;congruence. -left;reflexivity. -Defined. - -Theorem pos_eq_dec_refl : forall m, pos_eq_dec m m = left _ (refl_equal m). -fix 1;intros [mm|mm|]. -simpl; rewrite pos_eq_dec_refl; reflexivity. -simpl; rewrite pos_eq_dec_refl; reflexivity. -reflexivity. -Qed. - -Theorem pos_eq_dec_ex : forall m n, - pos_eq m n =true -> exists h:m=n, - pos_eq_dec m n = left _ h. -fix 1;intros [mm|mm|] [nn|nn|];try (simpl;congruence). -simpl;intro e. -elim (pos_eq_dec_ex _ _ e). -intros x ex; rewrite ex. -exists (f_equal xI x). -reflexivity. -simpl;intro e. -elim (pos_eq_dec_ex _ _ e). -intros x ex; rewrite ex. -exists (f_equal xO x). -reflexivity. -simpl. -exists (refl_equal xH). -reflexivity. -Qed. - -Fixpoint nat_eq (m n:nat) {struct m}: bool:= -match m, n with -O,O => true -| S mm,S nn => nat_eq mm nn -| _,_ => false -end. - -Theorem nat_eq_refl : forall m n, nat_eq m n = true -> m = n. -induction m;simpl;intro n;destruct n;congruence || -(intro e;apply f_equal;auto). -Defined. - -Theorem refl_nat_eq : forall n, nat_eq n n = true. -induction n;simpl;trivial. -Defined. - Fixpoint Lget (A:Set) (n:nat) (l:list A) {struct l}:option A := match l with nil => None | x::q => @@ -178,7 +59,8 @@ simpl;auto. Qed. Lemma Lget_app : forall (A:Set) (a:A) l i, -Lget i (l ++ a :: nil) = if nat_eq i (length l) then Some a else Lget i l. +Lget i (l ++ a :: nil) = if nat_beq i (length l) then Some a else Lget i l. +Proof. induction l;simpl Lget;simpl length. intros [ | i];simpl;reflexivity. intros [ | i];simpl. @@ -274,17 +156,20 @@ Qed. Theorem Tget_Tadd: forall i j a T, Tget i (Tadd j a T) = - match (i ?= j) Eq with + match (i ?= j) with Eq => PSome a | Lt => Tget i T | Gt => Tget i T end. +Proof. intros i j. -caseq ((i ?= j) Eq). -intro H;rewrite (Pcompare_Eq_eq _ _ H);intros a;clear i H. +case_eq (i ?= j). +intro H;rewrite (Pos.compare_eq _ _ H);intros a;clear i H. induction j;destruct T;simpl;try (apply IHj);congruence. +unfold Pos.compare. generalize i;clear i;induction j;destruct T;simpl in H|-*; destruct i;simpl;try rewrite (IHj _ H);try (destruct i;simpl;congruence);reflexivity|| congruence. +unfold Pos.compare. generalize i;clear i;induction j;destruct T;simpl in H|-*; destruct i;simpl;try rewrite (IHj _ H);try (destruct i;simpl;congruence);reflexivity|| congruence. Qed. @@ -308,7 +193,8 @@ Inductive Full : Store -> Type:= | F_push : forall a S, Full S -> Full (push a S). Theorem get_Full_Gt : forall S, Full S -> - forall i, (i ?= index S) Eq = Gt -> get i S = PNone. + forall i, (i ?= index S) = Gt -> get i S = PNone. +Proof. intros S W;induction W. unfold empty,index,get,contents;intros;apply Tget_Tempty. unfold index,get,push;simpl contents. @@ -335,16 +221,17 @@ Qed. Theorem get_push_Full : forall i a S, Full S -> get i (push a S) = - match (i ?= index S) Eq with + match (i ?= index S) with Eq => PSome a | Lt => get i S | Gt => PNone end. +Proof. intros i a S F. -caseq ((i ?= index S) Eq). -intro e;rewrite (Pcompare_Eq_eq _ _ e). +case_eq (i ?= index S). +intro e;rewrite (Pos.compare_eq _ _ e). destruct S;unfold get,push,index;simpl contents;rewrite Tget_Tadd. -rewrite Pcompare_refl;reflexivity. +rewrite Pos.compare_refl;reflexivity. intros;destruct S;unfold get,push,index;simpl contents;rewrite Tget_Tadd. simpl index in H;rewrite H;reflexivity. intro H;generalize H;clear H. @@ -357,8 +244,9 @@ Qed. Lemma Full_push_compat : forall i a S, Full S -> forall x, get i S = PSome x -> get i (push a S) = PSome x. +Proof. intros i a S F x H. -caseq ((i ?= index S) Eq);intro test. +case_eq (i ?= index S);intro test. rewrite (Pcompare_Eq_eq _ _ test) in H. rewrite (get_Full_Eq _ F) in H;congruence. rewrite <- H. @@ -391,7 +279,7 @@ get i S = PSome x -> In x S F. induction F. intro i;rewrite get_empty; congruence. intro i;rewrite get_push_Full;trivial. -caseq ((i ?= index S) Eq);simpl. +case_eq (i ?= index S);simpl. left;congruence. right;eauto. congruence. diff --git a/plugins/rtauto/Rtauto.v b/plugins/rtauto/Rtauto.v index f5d452f7e..8b5e42447 100644 --- a/plugins/rtauto/Rtauto.v +++ b/plugins/rtauto/Rtauto.v @@ -13,7 +13,6 @@ Require Import Bool. Declare ML Module "rtauto_plugin". -Ltac caseq t := generalize (refl_equal t); pattern t at -1; case t. Ltac clean:=try (simpl;congruence). Inductive form:Set:= @@ -66,11 +65,11 @@ end. Theorem form_eq_refl: forall p q, form_eq p q = true -> p = q. induction p;destruct q;simpl;clean. intro h;generalize (pos_eq_refl _ _ h);congruence. -caseq (form_eq p1 q1);clean. +case_eq (form_eq p1 q1);clean. intros e1 e2;generalize (IHp1 _ e1) (IHp2 _ e2);congruence. -caseq (form_eq p1 q1);clean. +case_eq (form_eq p1 q1);clean. intros e1 e2;generalize (IHp1 _ e1) (IHp2 _ e2);congruence. -caseq (form_eq p1 q1);clean. +case_eq (form_eq p1 q1);clean. intros e1 e2;generalize (IHp1 _ e1) (IHp2 _ e2);congruence. Qed. @@ -260,7 +259,7 @@ induction p;intros hyps F gl. (* cas Axiom *) Focus 1. -simpl;caseq (get p hyps);clean. +simpl;case_eq (get p hyps);clean. intros f nth_f e;rewrite <- (form_eq_refl e). apply project with p;trivial. @@ -273,10 +272,10 @@ apply IHp;try constructor;trivial. (* Cas Arrow_Elim *) Focus 1. -simpl check_proof;caseq (get p hyps);clean. -intros f ef;caseq (get p0 hyps);clean. +simpl check_proof;case_eq (get p hyps);clean. +intros f ef;case_eq (get p0 hyps);clean. intros f0 ef0;destruct f0;clean. -caseq (form_eq f f0_1);clean. +case_eq (form_eq f f0_1);clean. simpl;intros e check_p1. generalize (project F ef) (project F ef0) (IHp (hyps \ f0_2) (F_push f0_2 hyps F) gl check_p1); @@ -288,10 +287,10 @@ auto. (* cas Arrow_Destruct *) Focus 1. -simpl;caseq (get p1 hyps);clean. +simpl;case_eq (get p1 hyps);clean. intros f ef;destruct f;clean. destruct f1;clean. -caseq (check_proof (hyps \ f1_2 =>> f2 \ f1_1) f1_2 p2);clean. +case_eq (check_proof (hyps \ f1_2 =>> f2 \ f1_1) f1_2 p2);clean. intros check_p1 check_p2. generalize (project F ef) (IHp1 (hyps \ f1_2 =>> f2 \ f1_1) @@ -302,7 +301,7 @@ simpl;apply compose3;auto. (* Cas False_Elim *) Focus 1. -simpl;caseq (get p hyps);clean. +simpl;case_eq (get p hyps);clean. intros f ef;destruct f;clean. intros _; generalize (project F ef). apply compose1;apply False_ind. @@ -310,13 +309,13 @@ apply compose1;apply False_ind. (* Cas And_Intro *) Focus 1. simpl;destruct gl;clean. -caseq (check_proof hyps gl1 p1);clean. +case_eq (check_proof hyps gl1 p1);clean. intros Hp1 Hp2;generalize (IHp1 hyps F gl1 Hp1) (IHp2 hyps F gl2 Hp2). apply compose2 ;simpl;auto. (* cas And_Elim *) Focus 1. -simpl;caseq (get p hyps);clean. +simpl;case_eq (get p hyps);clean. intros f ef;destruct f;clean. intro check_p;generalize (project F ef) (IHp (hyps \ f1 \ f2) (F_push f2 (hyps \ f1) (F_push f1 hyps F)) gl check_p). @@ -324,7 +323,7 @@ simpl;apply compose2;intros [h1 h2];auto. (* cas And_Destruct *) Focus 1. -simpl;caseq (get p hyps);clean. +simpl;case_eq (get p hyps);clean. intros f ef;destruct f;clean. destruct f1;clean. intro H;generalize (project F ef) @@ -346,9 +345,9 @@ apply compose1;simpl;auto. (* cas Or_elim *) Focus 1. -simpl;caseq (get p1 hyps);clean. +simpl;case_eq (get p1 hyps);clean. intros f ef;destruct f;clean. -caseq (check_proof (hyps \ f1) gl p2);clean. +case_eq (check_proof (hyps \ f1) gl p2);clean. intros check_p1 check_p2;generalize (project F ef) (IHp1 (hyps \ f1) (F_push f1 hyps F) gl check_p1) (IHp2 (hyps \ f2) (F_push f2 hyps F) gl check_p2); @@ -356,7 +355,7 @@ simpl;apply compose3;simpl;intro h;destruct h;auto. (* cas Or_Destruct *) Focus 1. -simpl;caseq (get p hyps);clean. +simpl;case_eq (get p hyps);clean. intros f ef;destruct f;clean. destruct f1;clean. intro check_p0;generalize (project F ef) @@ -367,7 +366,7 @@ apply compose2;auto. (* cas Cut *) Focus 1. -simpl;caseq (check_proof hyps f p1);clean. +simpl;case_eq (check_proof hyps f p1);clean. intros check_p1 check_p2; generalize (IHp1 hyps F f check_p1) (IHp2 (hyps\f) (F_push f hyps F) gl check_p2); @@ -375,7 +374,7 @@ simpl; apply compose2;auto. Qed. Theorem Reflect: forall gl prf, if check_proof empty gl prf then [[gl]] else True. -intros gl prf;caseq (check_proof empty gl prf);intro check_prf. +intros gl prf;case_eq (check_proof empty gl prf);intro check_prf. change (interp_ctx empty F_empty [[gl]]) ; apply interp_proof with prf;assumption. trivial. -- cgit v1.2.3