Skip to content

Commit

Permalink
Updated for Coq 8.5 (merged in namespace)
Browse files Browse the repository at this point in the history
  • Loading branch information
antalsz committed Jun 15, 2016
2 parents f571c1a + e417b18 commit 685d392
Show file tree
Hide file tree
Showing 20 changed files with 94 additions and 69 deletions.
3 changes: 3 additions & 0 deletions .gitignore
Original file line number Diff line number Diff line change
@@ -1,3 +1,6 @@
/*.vo
/*.glob
/.*.aux
/*.v.d
CoqSrc.mk
_CoqProject
8 changes: 4 additions & 4 deletions AssocList.v
Original file line number Diff line number Diff line change
Expand Up @@ -20,9 +20,9 @@ Require Import Coq.FSets.FSets.
Require Import Coq.Lists.List.
Require Import Coq.Logic.Decidable.

Require Import CoqFSetDecide.
Require Import CoqListFacts.
Require Import LibTactics.
Require Import Metalib.CoqFSetDecide.
Require Import Metalib.CoqListFacts.
Require Import Metalib.LibTactics.


(* *********************************************************************** *)
Expand Down Expand Up @@ -635,7 +635,7 @@ Section BindsProperties.
binds x a (y ~ b) ->
x = y.
Proof.
clear. intros H. inversion H as [HEq | HIn].
clear. intros H. inversion H as [HEq | HIn].
inversion HEq; intuition.
inversion HIn.
Qed.
Expand Down
35 changes: 17 additions & 18 deletions AssumeList.v
Original file line number Diff line number Diff line change
Expand Up @@ -14,13 +14,13 @@ Require Import Coq.FSets.FSets.
Require Import Coq.Lists.List.
Require Import Coq.Logic.Decidable.

Require Import CoqFSetDecide.
Require Import CoqListFacts.
Require Import LibTactics.
Require Import Metalib.CoqFSetDecide.
Require Import Metalib.CoqListFacts.
Require Import Metalib.LibTactics.

Require Import MetatheoryAtom.
Require Import Metalib.MetatheoryAtom.
Import AtomSetImpl.
Require Import LibTactics.
Require Import Metalib.LibTactics.


(* *********************************************************************** *)
Expand Down Expand Up @@ -65,7 +65,7 @@ Require Import LibTactics.
.

Definition map (f1 : A -> A) (f2: B -> B) (E : list (asn A B)) : list (asn A B) :=
List.map (fun x => match x with
List.map (fun x => match x with
| VarAsn _ x a => VarAsn B x (f1 a)
| AltAsn _ b => AltAsn A (f2 b)
end) E.
Expand Down Expand Up @@ -96,17 +96,17 @@ Require Import LibTactics.
uniq E ->
x `notin` dom E ->
uniq ((x ~~ a) ++ E)
| uniq_alt : forall b E,
uniq E ->
| uniq_alt : forall b E,
uniq E ->
uniq (one (AltAsn A b) ++ E).

Fixpoint erase (E : list (asn A B)) : list B :=
match E with
match E with
| nil => nil
| (AltAsn _ b) :: E => b :: erase E
| (VarAsn _ x a) :: E => erase E
| (VarAsn _ x a) :: E => erase E
end.

Unset Implicit Arguments.
End Definitions.

Expand Down Expand Up @@ -462,7 +462,7 @@ Require Import LibTactics.
uniq (E ++ F) -> uniq E.
Proof.
clear. intros H; induction E as [ | a E']; simpl_asnlist in *.
apply uniq_nil.
apply uniq_nil.
destruct a.
apply uniq_push; inversion H; subst.
auto.
Expand Down Expand Up @@ -507,7 +507,7 @@ Require Import LibTactics.
intuition.
rewrite disjoint_one_l in Hd. simpl_asnlist. fsetdec.
apply uniq_alt.
inversion HE; subst.
inversion HE; subst.
intuition.
Qed.

Expand Down Expand Up @@ -870,9 +870,9 @@ Require Import LibTactics.
unfold binds in *. simpl. intros [? | ?].
destruct a0.
inversion H; subst.
left. congruence.
left. congruence.
inversion H.
destruct a0.
destruct a0.
right. auto.
right. auto.
Qed.
Expand All @@ -889,8 +889,8 @@ Require Import LibTactics.
clear. induction E as [ | y ? F ]; intros J; simpl_asnlist.
analyze_binds J.
analyze_binds J; subst; auto with set.
inversion H0;subst. simpl. fsetdec.
inversion H.
inversion H0;subst. simpl. fsetdec.
inversion H.
Qed.

Lemma fresh_app_l :
Expand All @@ -917,4 +917,3 @@ Require Import LibTactics.




10 changes: 5 additions & 5 deletions Attic/MetatheoryEnv.v
Original file line number Diff line number Diff line change
Expand Up @@ -19,13 +19,13 @@ Require Import Coq.FSets.FSets.
Require Import Coq.Lists.List.
Require Import Coq.Logic.Decidable.

Require Import CoqFSetDecide.
Require Import CoqListFacts.
Require Import Metalib.CoqFSetDecide.
Require Import Metalib.CoqListFacts.

Require Import AssocList.
Require Import MetatheoryAtom.
Require Import Metalib.AssocList.
Require Import Metalib.MetatheoryAtom.
Import AtomSetImpl.
Require Import LibTactics.
Require Import Metalib.LibTactics.


(* *********************************************************************** *)
Expand Down
5 changes: 5 additions & 0 deletions Attic/MyAssumeList.v
Original file line number Diff line number Diff line change
Expand Up @@ -126,6 +126,8 @@ Open Scope list_scope.
(** [dom] computes the domain of an assumption list, i.e., the
set consisting of its keys. *)

(* FIXME: commented out by simplicity, but the code should be fixed *)
(*
Fixpoint dom
(A B : Type) (E : list (asn A B))
: KeySet.t :=
Expand Down Expand Up @@ -1552,4 +1554,7 @@ Hint Resolve @binds_weaken.
Hint Immediate @binds_remove_mid @binds_In.
*)

End Make.

8 changes: 4 additions & 4 deletions CoqListFacts.v
Original file line number Diff line number Diff line change
Expand Up @@ -10,7 +10,7 @@
Require Import Coq.Lists.List.
Require Import Coq.Lists.SetoidList.

Require Import CoqUniquenessTac.
Require Import Metalib.CoqUniquenessTac.

Open Scope list_scope.

Expand Down Expand Up @@ -126,7 +126,7 @@ Qed.

Lemma InA_iff_In : forall (A : Type) (x : A) (xs : list A),
InA (@eq _) x xs <-> In x xs.
Proof.
Proof.
split; auto using InA_In.
apply SetoidList.In_InA. apply eq_equivalence.
Qed.
Expand Down Expand Up @@ -230,14 +230,14 @@ Section Uniqueness_Of_SetoidList_Proofs.

Theorem lelistA_unique :
forall (x : A) (xs : list A) (p q : lelistA R x xs), p = q.
Proof. induction p using lelistA_ind'; uniqueness 1. Qed.
Proof. induction p using lelistA_ind'; uniqueness 1. Qed.

Theorem sort_unique :
forall (xs : list A) (p q : sort R xs), p = q.
Proof. induction p using sort_ind'; uniqueness 1. apply lelistA_unique. Qed.

Theorem eqlistA_unique :
forall (xs ys : list A) (p q : eqlistA R xs ys), p = q.
Proof. induction p using eqlistA_ind'; uniqueness 2. Qed.
Proof. induction p using eqlistA_ind'; uniqueness 2. Qed.

End Uniqueness_Of_SetoidList_Proofs.
8 changes: 4 additions & 4 deletions CoqUniquenessTacEx.v
Original file line number Diff line number Diff line change
Expand Up @@ -9,7 +9,7 @@ Require Import Coq.Arith.Peano_dec.
Require Import Coq.Lists.SetoidList.
Require Import Coq.omega.Omega.

Require Import CoqUniquenessTac.
Require Import Metalib.CoqUniquenessTac.


(* *********************************************************************** *)
Expand All @@ -31,8 +31,8 @@ Proof.
induction p using le_ind';
uniqueness 1;
assert False by omega; intuition.
Qed.

Qed.


(* ********************************************************************** *)
Expand Down Expand Up @@ -64,7 +64,7 @@ Section Uniqueness_Of_SetoidList_Proofs.

Theorem eqlistA_unique :
forall (xs ys : list A) (p q : eqlistA R xs ys), p = q.
Proof. induction p using eqlistA_ind'; uniqueness 2. Qed.
Proof. induction p using eqlistA_ind'; uniqueness 2. Qed.

End Uniqueness_Of_SetoidList_Proofs.

Expand Down
2 changes: 1 addition & 1 deletion FSetExtra.v
Original file line number Diff line number Diff line change
Expand Up @@ -8,7 +8,7 @@

Require Import Coq.FSets.FSets.

Require Import CoqFSetInterface.
Require Import Metalib.CoqFSetInterface.


(* *********************************************************************** *)
Expand Down
2 changes: 1 addition & 1 deletion FSetWeakNotin.v
Original file line number Diff line number Diff line change
Expand Up @@ -14,7 +14,7 @@

Require Import Coq.FSets.FSetInterface.

Require Import CoqFSetDecide.
Require Import Metalib.CoqFSetDecide.


(* *********************************************************************** *)
Expand Down
2 changes: 1 addition & 1 deletion Fsub_LetSum_Definitions.v
Original file line number Diff line number Diff line change
Expand Up @@ -16,7 +16,7 @@
- #<a href="##auto">Automation</a>#
*)

Require Export Metatheory.
Require Export Metalib.Metatheory.


(* ********************************************************************** *)
Expand Down
2 changes: 1 addition & 1 deletion Fsub_LetSum_Infrastructure.v
Original file line number Diff line number Diff line change
Expand Up @@ -19,7 +19,7 @@
- #<a href="##auto">Automation</a>#
- #<a href="##body">Properties of body_e</a># *)

Require Export Fsub_LetSum_Definitions.
Require Export Metalib.Fsub_LetSum_Definitions.


(* ********************************************************************** *)
Expand Down
2 changes: 1 addition & 1 deletion Fsub_LetSum_Lemmas.v
Original file line number Diff line number Diff line change
Expand Up @@ -21,7 +21,7 @@
- #<a href="##regularity">Regularity lemmas</a>#
- #<a href="##auto">Automation</a># *)

Require Export Fsub_LetSum_Infrastructure.
Require Export Metalib.Fsub_LetSum_Infrastructure.


(* ********************************************************************** *)
Expand Down
2 changes: 1 addition & 1 deletion Fsub_LetSum_Soundness.v
Original file line number Diff line number Diff line change
Expand Up @@ -12,7 +12,7 @@
- #<a href="##preservation">Preservation</a>#
- #<a href="##progress">Progress</a># *)

Require Export Fsub_LetSum_Lemmas.
Require Export Metalib.Fsub_LetSum_Lemmas.


(* ********************************************************************** *)
Expand Down
4 changes: 2 additions & 2 deletions LibLNgen.v
Original file line number Diff line number Diff line change
Expand Up @@ -7,8 +7,8 @@

(** A library of code for supporting LNgen. *)

Require Export LibDefaultSimp.
Require Import Metatheory.
Require Export Metalib.LibDefaultSimp.
Require Import Metalib.Metatheory.
Require Import Omega.


Expand Down
24 changes: 21 additions & 3 deletions Makefile
Original file line number Diff line number Diff line change
Expand Up @@ -25,6 +25,15 @@ COQC = coqc
COQDEP = coqdep
COQDOC = coqdoc


## Library name used for the imports in Coq

LIBNAME=Metalib

## Name of the submakefile generated by coq_makefile

COQMKFILENAME=CoqSrc.mk

## Include directories, one per line.

INCDIRS = \
Expand All @@ -35,6 +44,7 @@ INCDIRS = \
DOCDIR = doc/html

## List of files to be compiled and documented.
# TODO: should be useless once we get coq_makefile to compile the documentation as well

FILES = \
AssocList \
Expand Down Expand Up @@ -75,10 +85,17 @@ INCFLAGS = $(foreach i, $(INCDIRS), -I $(i))
.PHONY: all clean coq dist doc documentation gens
.SUFFIXES: .v .vo

all: Metatheory.vo MetatheoryAlt.vo LibLNgen.vo
all: coq # Metatheory.vo MetatheoryAlt.vo LibLNgen.vo

#coq: $(VOFILES)
coq: $(COQMKFILENAME)
# FIXME: loose integration right now, but when the doc works we should be able to simply include the generated makefile
@$(MAKE) -f CoqSrc.mk

coq: $(VOFILES)
$(COQMKFILENAME): Makefile
{ echo "-R . $(LIBNAME) " ; ls *.v ; } > _CoqProject && coq_makefile -f _CoqProject -o $(COQMKFILENAME)

# TODO: in theory, coq_makefile creates targets for documentation, so we should be able to use it instead of handwritten rules
doc:
+make documentation

Expand All @@ -87,7 +104,8 @@ documentation: $(DOCDIR) $(VOFILES)
cp -f custom.css $(DOCDIR)/coqdoc.css

clean:
rm -f *.vo *.glob *.cmi *.cmx *.o
#rm -f *.vo *.glob *.cmi *.cmx *.o
@$(MAKE) -f CoqSrc.mk clean
rm -rf $(DOCDIR)

############################################################################
Expand Down
12 changes: 6 additions & 6 deletions Metatheory.v
Original file line number Diff line number Diff line change
Expand Up @@ -10,11 +10,11 @@ Require Export Coq.Arith.Arith.
Require Export Coq.FSets.FSets.
Require Export Coq.Lists.List.

Require Export AssocList.
Require Export CoqEqDec.
Require Export CoqListFacts.
Require Export LibTactics.
Require Export MetatheoryAtom.
Require Export Metalib.AssocList.
Require Export Metalib.CoqEqDec.
Require Export Metalib.CoqListFacts.
Require Export Metalib.LibTactics.
Require Export Metalib.MetatheoryAtom.


(* ********************************************************************** *)
Expand Down Expand Up @@ -292,5 +292,5 @@ Ltac apply_fresh_base H gather_vars atom_name :=

(* SCW added this one for list support *)
Set Implicit Arguments.
Definition union_map (A:Set) (f:A -> vars) (l:list A) :=
Definition union_map (A:Set) (f:A -> vars) (l:list A) :=
(List.fold_right (fun t acc => f t \u acc) {}) l.
12 changes: 6 additions & 6 deletions MetatheoryAlt.v
Original file line number Diff line number Diff line change
Expand Up @@ -10,13 +10,13 @@ Require Export Coq.Arith.Arith.
Require Export Coq.FSets.FSets.
Require Export Coq.Lists.List.

Require Export CoqEqDec.
Require Export CoqListFacts.
Require Export LibTactics.
Require Export MetatheoryAtom.
Require Export Metalib.CoqEqDec.
Require Export Metalib.CoqListFacts.
Require Export Metalib.LibTactics.
Require Export Metalib.MetatheoryAtom.
(* Require Export MetatheoryEnv. *)

Require Export AssumeList.
Require Export Metalib.AssumeList.

(* ********************************************************************** *)
(** * Decidable equality *)
Expand Down Expand Up @@ -94,7 +94,7 @@ Open Scope set_hs_scope.
standard recursive notation for lists, e.g., the one found in the
Program library of Coq's standard library. *)

(* Notation "[ x ]" := (EnvImpl.one x) : env_scope.
(* Notation "[ x ]" := (EnvImpl.one x) : env_scope.
Open Scope env_scope.
*)
Expand Down
Loading

0 comments on commit 685d392

Please sign in to comment.