pax_global_header00006660000000000000000000000064146057605110014517gustar00rootroot0000000000000052 comment=2a8e12360cceee510f39e3ef4d0a7472d70fa684 math-classes-8.19.0/000077500000000000000000000000001460576051100142025ustar00rootroot00000000000000math-classes-8.19.0/.github/000077500000000000000000000000001460576051100155425ustar00rootroot00000000000000math-classes-8.19.0/.github/workflows/000077500000000000000000000000001460576051100175775ustar00rootroot00000000000000math-classes-8.19.0/.github/workflows/docker-action.yml000066400000000000000000000021311460576051100230410ustar00rootroot00000000000000# This file was generated from `meta.yml`, please do not edit manually. # Follow the instructions on https://github.com/coq-community/templates to regenerate. name: Docker CI on: push: branches: - master pull_request: branches: - '**' jobs: build: # the OS must be GNU/Linux to be able to use the docker-coq-action runs-on: ubuntu-latest strategy: matrix: image: - 'coqorg/coq:dev' - 'coqorg/coq:8.19' - 'coqorg/coq:8.18' fail-fast: false steps: - uses: actions/checkout@v3 - uses: coq-community/docker-coq-action@v1 with: opam_file: 'coq-math-classes.opam' custom_image: ${{ matrix.image }} before_install: | startGroup "Setup and print opam config" opam repo -a --set-default add coq-extra-dev https://coq.inria.fr/opam/extra-dev opam config list; opam repo list; opam list endGroup # See also: # https://github.com/coq-community/docker-coq-action#readme # https://github.com/erikmd/docker-coq-github-action-demo math-classes-8.19.0/.gitignore000066400000000000000000000003441460576051100161730ustar00rootroot00000000000000_CoqProject Makefile.conf Makefile.bak Makefile.coq Makefile.coq.conf *.d *.vo *.vok *.vos *.glob *.pdf *.crashcoqide *.pyc tmp.* coqidescript coqdoc *# *~ *.ml *.cmi *.cmx *.cmxs *.o *.native *.aux result .lia.cache .nia.cache math-classes-8.19.0/LICENSE000066400000000000000000000021101460576051100152010ustar00rootroot00000000000000The MIT License (MIT) Copyright (c) 2009-present, Math Classes authors Permission is hereby granted, free of charge, to any person obtaining a copy of this software and associated documentation files (the "Software"), to deal in the Software without restriction, including without limitation the rights to use, copy, modify, merge, publish, distribute, sublicense, and/or sell copies of the Software, and to permit persons to whom the Software is furnished to do so, subject to the following conditions: The above copyright notice and this permission notice shall be included in all copies or substantial portions of the Software. THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE SOFTWARE. math-classes-8.19.0/Makefile000066400000000000000000000004401460576051100156400ustar00rootroot00000000000000all: Makefile.coq +make -f Makefile.coq all clean: Makefile.coq +make -f Makefile.coq clean rm -f Makefile.coq Makefile.coq: _CoqProject $(COQBIN)coq_makefile -f _CoqProject -o Makefile.coq _CoqProject: ./configure.sh %: Makefile.coq +make -f Makefile.coq $@ .PHONY: all clean math-classes-8.19.0/README.md000066400000000000000000000076031460576051100154670ustar00rootroot00000000000000 # Math Classes [![Docker CI][docker-action-shield]][docker-action-link] [![Contributing][contributing-shield]][contributing-link] [![Code of Conduct][conduct-shield]][conduct-link] [![Zulip][zulip-shield]][zulip-link] [![DOI][doi-shield]][doi-link] [docker-action-shield]: https://github.com/coq-community/math-classes/actions/workflows/docker-action.yml/badge.svg?branch=master [docker-action-link]: https://github.com/coq-community/math-classes/actions/workflows/docker-action.yml [contributing-shield]: https://img.shields.io/badge/contributions-welcome-%23f7931e.svg [contributing-link]: https://github.com/coq-community/manifesto/blob/master/CONTRIBUTING.md [conduct-shield]: https://img.shields.io/badge/%E2%9D%A4-code%20of%20conduct-%23f15a24.svg [conduct-link]: https://github.com/coq-community/manifesto/blob/master/CODE_OF_CONDUCT.md [zulip-shield]: https://img.shields.io/badge/chat-on%20zulip-%23c1272d.svg [zulip-link]: https://coq.zulipchat.com/#narrow/stream/237663-coq-community-devs.20.26.20users [doi-shield]: https://zenodo.org/badge/DOI/10.1017/S0960129511000119.svg [doi-link]: https://doi.org/10.1017/S0960129511000119 Math classes is a library of abstract interfaces for mathematical structures, such as: * Algebraic hierarchy (groups, rings, fields, …) * Relations, orders, … * Categories, functors, universal algebra, … * Numbers: N, Z, Q, … * Operations, (shift, power, abs, …) It is heavily based on Coq’s new type classes in order to provide: structure inference, multiple inheritance/sharing, convenient algebraic manipulation (e.g. rewriting) and idiomatic use of notations. ## Meta - Author(s): - Eelis van der Weegen (initial) - Bas Spitters (initial) - Robbert Krebbers (initial) - Coq-community maintainer(s): - Bas Spitters ([**@spitters**](https://github.com/spitters)) - License: [MIT License](LICENSE) - Compatible Coq versions: Coq 8.18 or later (use releases for other Coq versions) - Additional dependencies: - [BigNums](https://github.com/coq/bignums) - Coq namespace: `MathClasses` - Related publication(s): - [Type Classes for Mathematics in Type Theory](https://arxiv.org/abs/1102.1323) doi:[10.1017/S0960129511000119](https://doi.org/10.1017/S0960129511000119) ## Building and installation instructions The easiest way to install the latest released version of Math Classes is via [OPAM](https://opam.ocaml.org/doc/Install.html): ```shell opam repo add coq-released https://coq.inria.fr/opam/released opam install coq-math-classes ``` To instead build and install manually, do: ``` shell git clone https://github.com/coq-community/math-classes.git cd math-classes ./configure.sh make # or make -j make install ``` ## Directory structure ### categories/ Proofs that certain structures form categories. ### functors/ ### interfaces/ Definitions of abstract interfaces/structures. ### implementations/ Definitions of concrete data structures and algorithms, and proofs that they are instances of certain structures (i.e. implement certain interfaces). ### misc/ Miscellaneous things. ### orders/ Theory about orders on different structures. ### quote/ Prototype implementation of type class based quoting. To be integrated. ### theory/ Proofs of properties of structures. ### varieties/ Proofs that certain structures are varieties, and translation to/from type classes dedicated to these structures (defined in interfaces/). The reason we treat categories and varieties differently from other structures (like groups and rings) is that they are like meta-interfaces whose implementations are not concrete data structures and algorithms but are themselves abstract structures. To be able to distinguish the various arrows, we recommend using a variable width font. math-classes-8.19.0/SConscript000066400000000000000000000007001460576051100162110ustar00rootroot00000000000000 import os, glob, string nodes = ['.'] dirs = [] vs = [] while nodes: node = nodes.pop() b = os.path.basename(node) if node.endswith('.v') and not b.endswith('_nobuild.v'): vs += [File(node)] if os.path.isdir(node) and b != 'broken': dirs += [node] nodes += glob.glob(node + '/*') env = DefaultEnvironment() vos = globs = [] for node in vs: vo, glob = env.Coq(node) vos += [vo] globs += [glob] Return('vs vos globs') math-classes-8.19.0/SConstruct000066400000000000000000000015171460576051100162400ustar00rootroot00000000000000import os env = DefaultEnvironment(ENV = os.environ, tools=['default', 'Coq']) (vs, vos, globs) = env.SConscript(dirs='.') env['COQFLAGS'] = Rs = ' -R . MathClasses ' Default('implementations', 'theory', 'categories', 'orders', 'varieties', 'misc', 'functors') env.CoqDoc(env.Dir('coqdoc'), vs, COQDOCFLAGS='-utf8 --toc -g --no-lib-name http://coq.inria.fr/library') # Todo: Do "patch --backup $TARGET/coqdoc.css ../tools/coqdoc.css.diff", including the dependency on the .diff file. # Note: The generated documentation is no good, because of Coq bug #2423. vs_string = ' '.join(map(str, vs)) os.system('coqdep ' + Rs + vs_string + ' > deps') ParseDepends('deps') open('coqidescript', 'w').write('#!/bin/sh\ncoqide ' + Rs + ' $@ \n') os.chmod('coqidescript', 0755) env.Command('deps.dot', [], '../tools/DepsToDot.hs < deps > $TARGET') math-classes-8.19.0/_CoqProject.in000066400000000000000000000001071460576051100167400ustar00rootroot00000000000000# Library name -R . MathClasses # Coq files (filled by the configure) math-classes-8.19.0/categories/000077500000000000000000000000001460576051100163275ustar00rootroot00000000000000math-classes-8.19.0/categories/JMcat.v000066400000000000000000000052201460576051100175130ustar00rootroot00000000000000Require MathClasses.misc.JMrelation. Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.functors MathClasses.theory.categories. Record Object := object { obj:> Type ; Arrows_inst: Arrows obj ; Equiv_inst: ∀ x y: obj, Equiv (x ⟶ y) ; CatId_inst: CatId obj ; CatComp_inst: CatComp obj ; Category_inst: Category obj }. Arguments object _ {Arrows_inst Equiv_inst CatId_inst CatComp_inst Category_inst}. #[global] Existing Instance Arrows_inst. #[global] Existing Instance Equiv_inst. #[global] Existing Instance CatId_inst. #[global] Existing Instance CatComp_inst. #[global] Existing Instance Category_inst. Record Arrow (x y: Object): Type := arrow { map_obj:> obj x → obj y ; Fmap_inst: Fmap map_obj ; Functor_inst: Functor map_obj _ }. Arguments arrow {x y} _ _ _. Arguments map_obj {x y} _ _. #[global] Existing Instance Fmap_inst. #[global] Existing Instance Functor_inst. #[global] Instance: Arrows Object := Arrow. (* Hint Extern 4 (Arrows Object) => exact Arrow: typeclass_instances. *) Section contents. Section more_arrows. Context (x y: Object). Global Instance e: Equiv (x ⟶ y) := λ a b, (∀ v, a v ≡ b v) ∧ (∀ `(f: v ⟶ w), JMrelation.R (=) (fmap a f) _ (=) (fmap b f)). Instance e_refl: Reflexive e. Proof. intro a. unfold e. intuition. apply JMrelation.reflexive, _. Qed. Instance e_sym: Symmetric e. Proof with intuition. unfold e. intros ?? [P Q]... apply JMrelation.symmetric... Qed. Instance e_trans: Transitive e. Proof with intuition. unfold e. intros a b c [P Q] [R S]... transitivity (b v)... apply JMrelation.transitive with _ (=) (fmap b f)... Qed. Global Instance: Setoid (x ⟶ y). Proof. split; apply _. Qed. End more_arrows. Global Instance: CatId Object := λ _, arrow id _ _. Global Program Instance: CatComp Object := λ _ _ _ x y, arrow (x ∘ y) _ _. Global Instance: ∀ x y z: Object, Proper ((=) ==> (=) ==> (=)) ((◎): (y ⟶ z) → (x ⟶ y) → (x ⟶ z)). Proof with intuition; try apply _. unfold equiv, e. intros x y z a b [P Q] c d [R S]. split; intros. change (a (c v) ≡ b (d v)). congruence. change (JMrelation.R (=) (fmap a (fmap c f)) _ (=) (fmap b (fmap d f))). apply JMrelation.transitive with _ (=) (fmap a (fmap d f))... specialize (S _ _ f). revert S. generalize (fmap c f) (fmap d f). repeat rewrite R. intros. apply JMrelation.relate. rewrite (JMrelation.unJM _ _ _ _ _ S)... (* <- uses K *) Qed. Global Instance: Category Object. Proof. repeat (split; try apply _); intuition; apply reflexivity. Qed. End contents. math-classes-8.19.0/categories/algebras.v000066400000000000000000000037051460576051100203030ustar00rootroot00000000000000(* Show that algebras with homomorphisms between them form a category. *) Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.theory.categories. Require MathClasses.categories.setoids MathClasses.categories.product. Record Object (sign: Signature) : Type := object { algebra_carriers:> sorts sign → Type ; algebra_equiv: ∀ a, Equiv (algebra_carriers a) ; algebra_ops: AlgebraOps sign algebra_carriers ; algebra_proof: Algebra sign algebra_carriers }. Arguments object _ _ {algebra_equiv algebra_ops algebra_proof}. (* Avoid Coq trying to apply algebra_equiv to find arbitrary Equiv instances *) #[global] Hint Extern 0 (Equiv (algebra_carriers _ _ _)) => eapply @algebra_equiv : typeclass_instances. #[global] Existing Instance algebra_ops. #[global] Existing Instance algebra_proof. Section contents. Variable sign: Signature. Global Instance: Arrows (Object sign) := λ X Y, sig (HomoMorphism sign X Y). Program Definition arrow `{Algebra sign A} `{Algebra sign B} f `{!HomoMorphism sign A B f}: object sign A ⟶ object sign B := f. Global Program Instance: CatId (Object sign) := λ _ _, id. Global Program Instance comp: CatComp (Object sign) := λ _ _ _ f g v, f v ∘ g v. Global Program Instance: ∀ x y: Object sign, Equiv (x ⟶ y) := λ _ _ x y, ∀ b, pointwise_relation _ (=) (x b) (y b). Global Instance: ∀ x y: Object sign, Setoid (x ⟶ y). Proof. constructor. repeat intro. reflexivity. intros ? ? E ? ?. symmetry. apply E. intros ? ? ? E F ? ?. rewrite (E _ _). apply F. Qed. Instance: `{Proper ((=) ==> (=) ==> (=)) (comp x y z)}. Proof. intros ? ? ? x0 ? E ? ? F ? ?. simpl. unfold compose. do 3 red in E, F. destruct (proj2_sig x0). rewrite F, E. reflexivity. Qed. Global Instance: Category (Object sign). Proof. constructor; try apply _; repeat intro; reflexivity. Qed. End contents. math-classes-8.19.0/categories/categories.v000066400000000000000000000164711460576051100206540ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.functors MathClasses.theory.categories. Record Object := object { obj:> Type ; Arrows_inst: Arrows obj ; Equiv_inst: ∀ x y: obj, Equiv (x ⟶ y) ; CatId_inst: CatId obj ; CatComp_inst: CatComp obj ; Category_inst: Category obj }. Arguments object _ {Arrows_inst Equiv_inst CatId_inst CatComp_inst Category_inst}. #[global] Existing Instance Arrows_inst. #[global] Hint Extern 0 (Equiv (_ ⟶ _)) => eapply @Equiv_inst : typeclass_instances. #[global] Existing Instance CatId_inst. #[global] Existing Instance CatComp_inst. #[global] Existing Instance Category_inst. (* Todo: Ask mattam for [Existing Instance foo bar bas.] *) (* The above would be much more elegantly written as Inductive Object := object (obj: Type) `{Category obj}. Unfortunately, that doesn't register the coercion and class instances. *) Record Arrow (x y: Object): Type := arrow { map_obj:> obj x → obj y ; Fmap_inst: Fmap map_obj ; Functor_inst: Functor map_obj _ }. Arguments arrow {x y} _ {Fmap_inst Functor_inst}. Arguments map_obj {x y} _ _. #[global] Existing Instance Fmap_inst. #[global] Existing Instance Functor_inst. #[global] Instance: Arrows Object := Arrow. (* Hint Extern 4 (Arrows Object) => exact Arrow: typeclass_instances. *) (* Matthieu is adding [Existing Instance (c: T).], which is nicer. *) Section contents. Section more_arrows. Context (x y: Object). Global Program Instance e: Equiv (x ⟶ y) := λ a b, exists X: ∀ v, isoT _ _, ∀ (p q: x) (r: p ⟶ q), fmap a r ◎ snd (X p) = snd (X q) ◎ fmap b r. Let e_refl: Reflexive e. Proof. intro a. exists (λ v, refl_arrows (a v)). intros ???. simpl. rewrite left_identity, right_identity. reflexivity. Qed. Program Let sym_arrows (a b: x → y) (v: x) (p: isoT (a v) (b v)): isoT (b v) (a v) := (snd p, fst p). Next Obligation. destruct p. simpl in *. firstorder. Qed. Let e_sym: Symmetric e. Proof. intros ?? [x1 H]. exists (λ v, sym_arrows _ _ _ (x1 v)). simpl. intros ???. pose proof (H p q r). destruct (x1 p), (x1 q). simpl in *. apply (arrows_between_isomorphic_objects _ _ _ _ _ _ _ _ _ _ u u0). assumption. Qed. (* todo: clean up *) Program Let trans_arrows (x0 y0 z: x → y) (v: x) (x1: sig (λ (p: (x0 v ⟶ y0 v) * _), uncurry iso_arrows p)) (x2: sig (λ (p: (y0 v ⟶ z v) * _), uncurry iso_arrows p)): (* todo: use isoT *) isoT (x0 v) (z v) := (fst x2 ◎ fst x1, snd x1 ◎ snd x2). Next Obligation. Proof with assumption. destruct H as [? H1], H0 as [? H2]. unfold uncurry. simpl in *. split. rewrite <- associativity, (associativity a1 a2 a0), H0, left_identity... rewrite <- associativity, (associativity a0 a a1), H1, left_identity... Qed. Let e_trans: Transitive e. Proof. intros a b c [f H] [g H0]. exists (λ v, trans_arrows _ _ _ _ (f v) (g v)). simpl. intros ? ? ?. generalize (H p q r), (H0 p q r). clear H H0. intros E E'. rewrite associativity, E, <- associativity, E', associativity. reflexivity. Qed. Global Instance: Setoid (x ⟶ y). Proof. split; assumption. Qed. End more_arrows. Instance obj_iso (x: Object): Equiv x := @iso x _ _ _ _. Typeclasses Transparent Arrows. Global Instance: ∀ (x y: Object) (a: x ⟶ y), Setoid_Morphism (map_obj a). Proof with try apply _. constructor... intros v w [[f g] [E F]]. exists (fmap a f, fmap a g). unfold uncurry. destruct a. simpl in *. split; rewrite <- preserves_comp... rewrite E. apply preserves_id... rewrite F. apply preserves_id... Qed. (* Putting this in the "arrows" section above (where it belongs) triggers a Coq bug. *) Global Instance: CatId Object := λ _, arrow id. Global Program Instance: CatComp Object := λ _ _ _ x y, arrow (x ∘ y). Program Let proper_arrows (x y z: Object) (x0 y0: y ⟶ z) (x1 y1: x ⟶ y) (f: ∀ v, isoT (map_obj x0 v) (map_obj y0 v)) (g: ∀ v, isoT (map_obj x1 v) (map_obj y1 v)) (v: x): isoT (map_obj x0 (map_obj x1 v)) (map_obj y0 (map_obj y1 v)) := (fst (f (y1 v)) ◎ fmap x0 (fst (g v)), fmap x0 (snd (g v)) ◎ snd (f (y1 v))). (* Todo: Investigate why things go wrong without the underscores. *) Next Obligation. Proof with try apply _; intuition. destruct (f (y1 v)) as [? [e0 e1]]. destruct (g v) as [? [e2 e3]]. split; simpl in *. rewrite <- associativity. rewrite (associativity (fmap x0 _) (fmap x0 _) _). rewrite <- preserves_comp, e2, preserves_id, left_identity... rewrite <- associativity. rewrite (associativity _ _ (fmap x0 _)). rewrite e1, left_identity, <- preserves_comp, e3, preserves_id... Qed. Global Instance: ∀ x y z: Object, Proper ((=) ==> (=) ==> (=)) ((◎): (y ⟶ z) → (x ⟶ y) → (x ⟶ z)). Proof with try apply _. repeat intro. unfold e. destruct H. destruct H0. simpl in *. exists (proper_arrows x y z x0 y0 x1 y1 x2 x3). intros. simpl. pose proof (H0 p q r). clear H0. destruct (x3 p) as [[a a0] [e0 e1]], (x3 q) as [[a1 a2] [e2 e3]]. clear x3. simpl in *. change ( fmap x0 (fmap x1 r) ◎ (fmap x0 a0 ◎ snd (` (x2 (y1 p)))) = fmap x0 a2 ◎ snd (` (x2 (y1 q))) ◎ fmap y0 (fmap y1 r)). pose proof (H (y1 p) (y1 q) (fmap y1 r)). clear H. destruct (x2 (y1 p)) as [[a3 a4] [e4 e5]], (x2 (y1 q)) as [[a5 a6] [e6 e7]]. clear x2. simpl in *. rewrite <- associativity, <- H0. clear H0. eapply (transitivity (y:=((fmap x0 (fmap x1 r) ◎ fmap x0 a0) ◎ a4))). repeat rewrite associativity. reflexivity. rewrite <- preserves_comp... rewrite H1. rewrite associativity. rewrite <- preserves_comp... reflexivity. Qed. (* todo: clean up! *) Program Let id_lr_arrows (x y: Object) (a: y ⟶ x) v: isoT (map_obj a v) (map_obj a v) := (cat_id, cat_id). (* We can't remove the map_obj here and elsewhere even though it's a coercion, because unification isn't smart enough to resolve and use that coercion. This is likely due to Coq bug #2229. *) Next Obligation. split; apply left_identity. Qed. Instance: ∀ x y: Object, LeftIdentity (comp x _ y) cat_id. Proof. intros ?? a. exists (id_lr_arrows _ _ a). intros ? ? ?. simpl. unfold compose, id. rewrite right_identity, left_identity. reflexivity. Qed. Instance: ∀ x y: Object, RightIdentity (comp x _ y) cat_id. Proof. intros ?? a. exists (id_lr_arrows _ _ a). intros ? ? ?. simpl. unfold compose, id. rewrite right_identity, left_identity. reflexivity. Qed. Section associativity. Variables (w x y z: Object) (a: w ⟶ x) (b: x ⟶ y) (c: y ⟶ z). Program Definition associativity_arrows (v: w): isoT (c (b (a v))) (c (b (a v))) := (fmap c (fmap b (fmap a cat_id)), fmap c (fmap b (fmap a cat_id))). Next Obligation. unfold uncurry. simpl. split; repeat rewrite preserves_id; try apply _; apply left_identity. Qed. End associativity. Instance: ArrowsAssociative Object. Proof. repeat intro. exists (associativity_arrows _ _ _ _ z0 y0 x0). simpl. intros ? ? ?. unfold compose. rewrite ! preserves_id; try apply _. (* todo: remove need for [try apply _] *) rewrite left_identity, right_identity. reflexivity. Qed. Global Instance: Category Object := {}. End contents. math-classes-8.19.0/categories/dual.v000066400000000000000000000041361460576051100174470ustar00rootroot00000000000000Require Import Coq.Relations.Relation_Definitions MathClasses.interfaces.abstract_algebra MathClasses.theory.categories MathClasses.interfaces.functors. Section contents. Context `{c: @Category Object A Aeq Aid Acomp}. Instance flipA: Arrows Object := flip A. Global Instance: @CatId Object flipA := Aid. Global Instance: @CatComp Object flipA := λ _ _ _, flip (Acomp _ _ _). Global Instance e: ∀ x y, Equiv (flipA x y) := λ x y, Aeq y x. Global Instance: ∀ (x y: Object), Setoid (flipA x y). Proof. intros. change (Setoid (A y x)). apply arrow_equiv. Qed. Instance: ∀ (x y z: Object), Proper ((=) ==> (=) ==> (=)) (@comp Object flipA _ x y z). Proof. intros x y z ? ? E ? ? F. change (Acomp z y x x1 x0 = Acomp z y x y1 y0). unfold equiv, e. destruct c. rewrite E, F. reflexivity. Qed. Global Instance cat: @Category Object flipA _ _ _. Proof with auto. destruct c. constructor; try apply _; auto. unfold comp, Arrow, flip. repeat intro. symmetry. apply comp_assoc. intros. apply id_r. intros. apply id_l. Qed. End contents. (** Avoid looping on applications of e, without making flipA opaque. Note that the dual instances make proof search slower in general, making flipA opaque for resolution would speed things up but require a few changes to the scripts to explicitly convert terms to applications of flipA. *) #[global] Hint Cut [_* e (_*) e] : typeclass_instances. Section functors. (** Given a functor F: C → D, we have a functor F^op: C^op → D^op *) Context {C D} F `{Functor C (H:=Ce) D (H1:=De) F}. Definition fmap_op: @Fmap _ flipA _ flipA F := fun v w => @fmap _ _ _ _ F _ w v. Global Instance: Functor F fmap_op. Proof with intuition. unfold e, fmap_op, flipA, flip, CatId_instance_0, CatComp_instance_0, flip. pose proof (functor_from F). pose proof (functor_to F). constructor; repeat intro. apply cat. apply cat. destruct (functor_morphism F b a). constructor... set (preserves_id F a)... apply (@preserves_comp _ _ Ce _ _ _ _ De _ _ F)... Qed. End functors. math-classes-8.19.0/categories/empty.v000066400000000000000000000013511460576051100176540ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.functors. Definition Empty_map {A: Empty_set → Type} : ∀ x : Empty_set, A x := λ x, match x with end. Local Notation E := Empty_map. #[global] Instance: Arrows Empty_set := E. #[global] Instance: CatComp Empty_set := E. #[global] Instance: CatId Empty_set := E. #[global] Instance: ∀ x y, Equiv (x ⟶ y) := E. #[global] Instance: ∀ x y, Setoid (x ⟶ y) := E. #[global] Instance: Category Empty_set. Proof. constructor; exact E. Qed. Section another_category. Context `{Category C}. Global Instance: Fmap (E: _ → C) := E. Global Instance: Functor (E: _ → C) E. Proof. constructor; exact E || typeclasses eauto. Qed. End another_category. math-classes-8.19.0/categories/functors.v000066400000000000000000000064641460576051100203730ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.functors MathClasses.theory.categories. Section natural_transformations_id_comp. Context `{Category A} `{Category B} `{!Functor (f: A → B) f'} `{!Functor (f: A → B) f'} `{!Functor (g: A → B) g'} `{!Functor (h: A → B) h'} `{!NaturalTransformation (m:f⇛g)} `{!NaturalTransformation (n:g⇛h)}. Global Instance id_transformation: NaturalTransformation (λ a, cat_id: f a ⟶ f a). Proof. intros ???. rewrite left_identity, right_identity; reflexivity. Qed. Global Instance compose_transformation: NaturalTransformation (λ a, n a ◎ m a). Proof. intros ???. rewrite <- associativity, natural, associativity, natural, associativity; reflexivity. Qed. End natural_transformations_id_comp. Record Object A `{Arrows A} `{∀ x y : A, Equiv (x ⟶ y)} `{!CatId A} `{!CatComp A} B `{Arrows B} `{∀ x y : B, Equiv (x ⟶ y)} `{!CatId B} `{!CatComp B} : Type := object { map_obj:> A → B ; Fmap_inst:> Fmap map_obj ; Functor_inst: Functor map_obj _ }. #[global] Existing Instance Fmap_inst. #[global] Existing Instance Functor_inst. Arguments Object _ {Aarrows Aeq Aid Acomp} _ {Barrows Beq Bid Bcomp} : rename. Arguments object {A Aarrows Aeq Aid Acomp B Barrows Beq Bid Bcomp} {Fmap Functor} _ : rename. Arguments map_obj {A Aarrows Aeq Aid Acomp B Barrows Beq Bid Bcomp} _ _ : rename. Record Arrow `{Arrows A} `{∀ x y : A, Equiv (x ⟶ y)} `{!CatId A} `{!CatComp A} `{Arrows B} `{∀ x y : B, Equiv (x ⟶ y)} `{!CatId B} `{!CatComp B} (F G : Object A B) : Type := arrow { eta:> map_obj F ⇛ map_obj G ; NaturalTransformation_inst: NaturalTransformation eta }. #[global] Existing Instance NaturalTransformation_inst. Arguments arrow {A Aarrows Aeq Aid Acomp B Barrows Beq Bid Bcomp F G} _ _ : rename. Arguments eta {A Aarrows Aeq Aid Acomp B Barrows Beq Bid Bcomp F G} _ _ : rename. Section contents. Context `{Category A} `{Category B}. Global Instance: Arrows (Object A B) := Arrow. Section arrow_setoid. Context (F G : Object A B). Global Program Instance e: Equiv (F ⟶ G) := λ m n, ∀ x: A, m x = n x. Instance e_refl: Reflexive e. Proof. intro a; red; reflexivity. Qed. Instance e_sym: Symmetric e. Proof. intros m n Hmn a. red in Hmn. rewrite Hmn. reflexivity. Qed. Instance e_trans: Transitive e. Proof. intros m n o Hmn Hno a. red in Hmn, Hno. rewrite Hmn, Hno. reflexivity. Qed. Instance: Equivalence e := {}. Global Instance: Setoid (F ⟶ G) := {}. End arrow_setoid. Global Instance: CatId (Object A B) := λ _, arrow (λ _, cat_id) _. Global Instance: CatComp (Object A B) := λ _ _ _ m n, arrow (λ a, m a ◎ n a) _. Global Instance: ∀ x y z: Object A B, Proper ((=) ==> (=) ==> (=)) ((◎): (y ⟶ z) → (x ⟶ y) → (x ⟶ z)). Proof. intros ????? Hx ?? Hy a. simpl. rewrite (Hx a), (Hy a). reflexivity. Qed. Instance: ∀ x y: Object A B, LeftIdentity (comp x y y) cat_id. Proof. repeat intro. simpl. apply left_identity. Qed. Instance: ∀ x y: Object A B, RightIdentity (comp x x y) cat_id. Proof. repeat intro. simpl. apply right_identity. Qed. Instance: ArrowsAssociative (Object A B). Proof. repeat intro. simpl. apply associativity. Qed. Global Instance: Category (Object A B) := {}. End contents. math-classes-8.19.0/categories/orders.v000066400000000000000000000040421460576051100200140ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.theory.categories MathClasses.orders.maps MathClasses.interfaces.orders MathClasses.orders.orders MathClasses.interfaces.functors. Require MathClasses.categories.setoids. Inductive Object := object { T:> Type; e: Equiv T; le: Le T; setoid_proof: Setoid T; po_proof: PartialOrder le }. #[global] Existing Instance e. #[global] Existing Instance le. #[global] Existing Instance setoid_proof. #[global] Existing Instance po_proof. Arguments object _ {e le setoid_proof po_proof}, _ {e} _ {setoid_proof po_proof}, _ _ _ {setoid_proof po_proof}. Section contents. Global Instance Arrow: Arrows Object := λ A B, sig (@OrderPreserving A B _ _ _ _). Global Program Instance: ∀ x y: Object, Equiv (x ⟶ y) := λ _ _, respectful (=) (=). Existing Instance order_morphism_mor. Global Instance: ∀ x y: Object, Setoid (x ⟶ y). Proof with intuition. intros x y. constructor. intros [? ?] ? ? E. now rewrite E. intros ? ? E ? ? ?. symmetry... intros [f Pf] [g Pg] [h Ph] E1 E2 a b E3. simpl. transitivity (g a)... Qed. Global Program Instance: CatId Object := λ _, id. Local Obligation Tactic := idtac. Global Program Instance: CatComp Object := λ _ _ _, compose. Instance: ∀ x y z: Object, Proper ((=) ==> (=) ==> (=)) (comp x y z). Proof. repeat intro. simpl. firstorder. Qed. Global Instance: Category Object. Proof. constructor; try apply _. intros ? ? ? ? [??] [??] [??] ? ? E. simpl. now rewrite E. intros ? ? [??] ? ? E. simpl. now rewrite E. intros ? ? [??] ? ? E. simpl. now rewrite E. Qed. Definition forget (O: Object) : setoids.Object := setoids.object O. Global Program Instance: Fmap forget := λ x y f, f. Global Instance: Functor forget _. Proof. constructor; try typeclasses eauto. * constructor; try typeclasses eauto. intros ???. assumption. * intros ????. assumption. * intros ? ? [] ? [] ? ? ?. simpl. rewrite H. reflexivity. Qed. End contents. math-classes-8.19.0/categories/product.v000066400000000000000000000122341460576051100202000ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra Coq.Logic.ChoiceFacts MathClasses.interfaces.functors MathClasses.theory.categories MathClasses.categories.categories. (* Axiom dependent_functional_choice: DependentFunctionalChoice. *) Section object. Context {I: Type} (O: I → Type) `{∀ i, Arrows (O i)} `{∀ i (x y: O i), Equiv (x ⟶ y)} `{∀ i, CatId (O i)} `{∀ i, CatComp (O i)} `{∀ i, Category (O i)}. Definition Object := ∀ i, O i. Global Instance pa: Arrows Object := λ x y, ∀ i, x i ⟶ y i. (* todo: make nameless *) Global Instance: CatId Object := λ _ _, cat_id. Global Instance: CatComp Object := λ _ _ _ d e i, d i ◎ e i. Definition e (x y: Object): Equiv (x ⟶ y) := λ f g, ∀ i, f i = g i. End object. (* Avoid Coq trying to apply e to find arbitrary Equiv instances *) #[global] Hint Extern 0 (Equiv (_ ⟶ _)) => eapply @e : typeclass_instances. Section contents. Context {I: Type} (O: I → Type) `{∀ i, Arrows (O i)} `{∀ i (x y: O i), Equiv (x ⟶ y)} `{∀ i, CatId (O i)} `{∀ i, CatComp (O i)} `{∀ i, Category (O i)}. Global Instance: ∀ x y: Object O, Setoid (x ⟶ y) := _. Global Instance: Category (Object O). Proof with try reflexivity. constructor. apply _. intros ? ? ? x y E x' y' F i. change (x i ◎ x' i = y i ◎ y' i). rewrite (E i), (F i)... repeat intro. apply comp_assoc. repeat intro. apply id_l. (* todo: Use left_identity *) repeat intro. apply id_r. Qed. Let product_object := categories.object (Object O). Notation ith_obj i := (categories.object (O i)). Program Definition project i: categories.object (Object O) ⟶ ith_obj i := @categories.arrow _ _ (λ d, d i) (λ _ _ a, a i) _. Next Obligation. Proof. (* functorial *) constructor; intros; try reflexivity; try apply _. constructor; try apply _. intros ? ? E. apply E. Qed. Section factors. Variables (C: categories.Object) (X: ∀ i, C ⟶ ith_obj i). Let ith_functor i := categories.Functor_inst _ _ (X i). (* todo: really necessary? *) Program Definition factor: C ⟶ product_object := @categories.arrow _ _ (λ (c: C) i, X i c) (λ (x y: C) (c: x ⟶ y) i, fmap (X i) c) _. Next Obligation. Proof with try reflexivity; intuition. (* functorial *) constructor; intros; try apply _. constructor; try apply _. intros ? ? E ?. change (fmap (X i) x = fmap (X i) y). rewrite E... intro. unfold fmap at 1. rewrite preserves_id... destruct X... intro. unfold fmap at 1. rewrite preserves_comp... destruct X... Qed. (* todo: those [destruct X]'s shouldn't be necessary *) (* Lemma s: is_sole (λ h', ∀ i, X i = project i ◎ h') factor. Proof with try reflexivity; intuition. split. intro. exists (λ v, refl_arrows (X i v)). simpl. unfold compose. intros ? ? ?. rewrite right_identity, left_identity... intros alt alt_factors. generalize (dependent_functional_choice I _ _ alt_factors). clear alt_factors. unfold isoT in *. simpl. intros [x H4]. unfold equiv. unfold categories.e. unfold compose in H4. set (P := λ v, prod (alt v ⟶ (λ i, (X i) v)) ((λ i, (X i) v) ⟶ alt v)). set (d := λ v, (λ i, snd (` (x i v)), λ i, fst (` (x i v))): P v). assert (∀ v, uncurry iso_arrows (d v)) as Q. split; simpl; intro. change (snd (` (x i v)) ◎ fst (` (x i v)) = cat_id). destruct (x i v) as [? []]... change (fst (` (x i v)) ◎ snd (` (x i v)) = cat_id). destruct (x i v) as [? []]... exists (λ v, exist (uncurry iso_arrows) _ (Q v)). intros p q r i. simpl. unfold comp. unfold CatComp_instance_0. (* todo: no! *) pose proof (H4 i p q r). clear H4. destruct (x i p) as [aa0 i0]. destruct (x i q) as [a1a2 i1]. simpl in *. unfold uncurry in *. unfold iso_arrows in *. destruct (categories.Functor_inst _ _ alt). unfold compose in x, aa0, a1a2. simpl in *. unfold fmap. match goal with |- appcontext [ categories.Fmap_inst _ ?cat ?alt ] => set (f:=categories.Fmap_inst _ cat alt) in |- * end. setoid_rewrite <- (left_identity (f p q r i ◎ fst aa0)). transitivity ((fst a1a2 ◎ snd a1a2) ◎ (f p q r i ◎ fst aa0)). apply comp_proper... apply transitivity with ((fst a1a2) ◎ (((snd a1a2) ◎ (categories.Fmap_inst _ _ alt p q r i)) ◎ (fst aa0))). repeat rewrite associativity... simpl. rewrite <- H5. repeat rewrite <- (associativity _ _). rewrite (proj2 i0), right_identity... Qed. (* WARNING: Uses DependentFunctionalChoice. (Todo: reflect.) *) (* todo: awful proof. clean up! *) *) End factors. (* At the time of this writing (Coq trunk 12801), attempting to state the above in terms of the nice product-relates interfaces from theory/categories results in universe inconsistency errors that I have isolated and reported as Coq bug #2252. *) Global Instance mono (X Y: Object O): ∀ (a: X ⟶ Y), (∀ i, @Mono _ _ (H0 _) (H2 i) _ _ (a i)) → Mono a. Proof. firstorder. Qed. (* todo: why so ugly all of a sudden? *) End contents. math-classes-8.19.0/categories/setoids.v000066400000000000000000000043421460576051100201730ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.theory.categories. Inductive Object := object { T:> Type; e: Equiv T; setoid_proof: Setoid T }. Arguments object _ {e setoid_proof}. #[global] Existing Instance e. #[global] Existing Instance setoid_proof. Section contents. Global Instance Arrow: Arrows Object := λ A B, sig (@Setoid_Morphism A B _ _). Global Program Instance: ∀ x y: Object, Equiv (x ⟶ y) := λ _ _, respectful (=) (=). Global Instance: ∀ x y: Object, Setoid (x ⟶ y). Proof with intuition. intros x y. constructor. intros [? ?] ? ? E. now rewrite E. intros ? ? E ? ? ?. symmetry... intros [f Pf] [g Pg] [h Ph] E1 E2 a b E3. simpl. transitivity (g a)... Qed. Global Program Instance: CatId Object := λ _, id. Local Obligation Tactic := idtac. Global Program Instance: CatComp Object := λ _ _ _, compose. Instance: ∀ x y z: Object, Proper ((=) ==> (=) ==> (=)) (comp x y z). Proof. repeat intro. simpl. firstorder. Qed. Global Instance: Category Object. Proof. constructor; try apply _. intros ? ? ? ? [??] [??] [??] ? ? E. simpl. now rewrite E. intros ? ? [??] ? ? E. simpl. now rewrite E. intros ? ? [??] ? ? E. simpl. now rewrite E. Qed. Global Instance: Producer Object := λ _ c, @object (∀ i, c i) (λ x y, ∀ i, x i = y i) _. (* todo: use pointwise_relation or something like that *) Section product. Context {Index: Type} (c: Index → Object). Global Program Instance: ElimProduct c (product c) := λ i x, x i. Next Obligation. constructor; try apply _. firstorder. Qed. Global Program Instance: IntroProduct c (product c) := λ d df x y, df y x. Next Obligation. constructor; try apply _. repeat intro. destruct df. simpl. firstorder. Qed. Global Instance: Product c (product c). Proof. split. intros ? ? ? E. simpl. unfold compose. destruct ccomp. rewrite E. reflexivity. intros ? E' ? x E i. simpl in *. symmetry in E |- *. apply (E' i _ _ E). Qed. End product. Global Instance: HasProducts Object := {}. Global Instance mono (X Y: Object) (a: X ⟶ Y): Injective (` a) → Mono a. Proof. intros A ??? E1 ?? E2. apply A. apply (E1 _ _ E2). Qed. End contents. math-classes-8.19.0/categories/unit.v000066400000000000000000000011671460576051100175020ustar00rootroot00000000000000Require Import Coq.Classes.RelationClasses Coq.Classes.Equivalence MathClasses.categories.categories MathClasses.interfaces.abstract_algebra MathClasses.categories.functors. #[global] Instance: Arrows unit := λ _ _, unit. #[global] Instance: CatId unit := λ _, tt. #[global] Instance: CatComp unit := λ _ _ _ _ _, tt. #[global] Instance: ∀ x y : unit, Equiv (x ⟶ y) := λ _ _, eq. #[global] Instance: ∀ x y : unit, Setoid (x ⟶ y) := {}. #[global] Instance: Category unit. Proof. constructor; try constructor; compute; repeat intro; repeat match goal with [ x : unit |- _ ] => destruct x end; reflexivity. Qed. math-classes-8.19.0/categories/varieties.v000066400000000000000000000037231460576051100205160ustar00rootroot00000000000000(* Show that Varieties with homomorphisms between them form a category. Hm, this file is almost identical to categories/algebra because the morphism are the same. There must be some way to factor out the commonality. *) Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms. Record Object (et: EquationalTheory) : Type := object { variety_carriers:> sorts et → Type ; variety_equiv: ∀ a, Equiv (variety_carriers a) ; variety_ops: AlgebraOps et variety_carriers ; variety_proof: InVariety et variety_carriers }. Arguments object _ _ {variety_equiv variety_ops variety_proof}. (* Avoid Coq trying to apply variety_equiv to find arbitrary Equiv instances *) #[global] Hint Extern 0 (Equiv (variety_carriers _ _ _)) => eapply @variety_equiv : typeclass_instances. #[global] Existing Instance variety_ops. #[global] Existing Instance variety_proof. Section contents. Variable et: EquationalTheory. Global Instance: Arrows (Object et) := λ X Y, sig (HomoMorphism et X Y). Program Definition arrow `{InVariety et A} `{InVariety et B} f `{!HomoMorphism et A B f}: object et A ⟶ object et B := f. Global Program Instance: CatId (Object et) := λ _ _, id. Global Program Instance: CatComp (Object et) := λ _ _ _ f g v, f v ∘ g v. Global Program Instance: ∀ (x y: Object et), Equiv (x ⟶ y) := λ _ _ x y, ∀ b, pointwise_relation _ (=) (x b) (y b). Global Instance: ∀ (x y: Object et), Setoid (x ⟶ y). Proof. constructor. repeat intro. reflexivity. intros ? ? E ? ?. symmetry. apply E. intros ? ? ? E F ? ?. rewrite (E _ _). apply F. Qed. Instance: ∀ (x y z: Object et), Proper ((=) ==> (=) ==> (=)) (comp x y z). Proof. intros ??? ? ? E ? ? F ? ?. simpl. unfold compose. rewrite (F _ _), (E _ _). reflexivity. Qed. Global Instance: Category (Object et). Proof. constructor; try apply _; repeat intro; reflexivity. Qed. End contents. math-classes-8.19.0/configure.sh000077500000000000000000000001601460576051100165170ustar00rootroot00000000000000#!/bin/sh cp -f _CoqProject.in _CoqProject find . -name "*.v" |grep -v misc/benchmarks_nobuild.v >> _CoqProject math-classes-8.19.0/coq-math-classes.opam000066400000000000000000000021631460576051100202260ustar00rootroot00000000000000opam-version: "2.0" maintainer: "b.a.w.spitters@gmail.com" version: "dev" homepage: "https://github.com/coq-community/math-classes" dev-repo: "git+https://github.com/coq-community/math-classes.git" bug-reports: "https://github.com/coq-community/math-classes/issues" license: "MIT" synopsis: "A library of abstract interfaces for mathematical structures in Coq" description: """ Math classes is a library of abstract interfaces for mathematical structures, such as: * Algebraic hierarchy (groups, rings, fields, …) * Relations, orders, … * Categories, functors, universal algebra, … * Numbers: N, Z, Q, … * Operations, (shift, power, abs, …) It is heavily based on Coq’s new type classes in order to provide: structure inference, multiple inheritance/sharing, convenient algebraic manipulation (e.g. rewriting) and idiomatic use of notations. """ build: [ [ "./configure.sh" ] [ make "-j%{jobs}%" ] ] install: [make "install"] depends: [ "coq" {(>= "8.18" & < "8.20~") | (= "dev")} "coq-bignums" ] tags: [ "logpath:MathClasses" ] authors: [ "Eelis van der Weegen" "Bas Spitters" "Robbert Krebbers" ] math-classes-8.19.0/functors/000077500000000000000000000000001460576051100160455ustar00rootroot00000000000000math-classes-8.19.0/functors/constant.v000066400000000000000000000021671460576051100200730ustar00rootroot00000000000000Require Import MathClasses.theory.categories MathClasses.interfaces.abstract_algebra MathClasses.interfaces.functors. Section constant_functor. Context `{Category A} `{Category B}. Context (b:B). Notation F := (const b: A → B). Global Instance: Fmap F := λ _ _ _, cat_id. Global Instance: ∀ a a' : A, Setoid_Morphism (fmap F : (a ⟶ a') → (F a ⟶ F a')). Proof. intros; constructor; try apply _. now intros ? ? ?. Qed. Global Instance ConstantFunctor : Functor F (λ _ _ _, cat_id). Proof. split; try apply _. reflexivity. intros; symmetry; compute; apply left_identity. Qed. End constant_functor. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Transparent const. Set Warnings "+unsupported-attributes". Section constant_transformation. Context `{Category A} `{Category J}. Context {a a' : A}. Global Instance constant_transformation {f : a⟶a'} : NaturalTransformation (const f : J → _). Proof. intros ? ? ?. now rewrite left_identity, right_identity. Qed. End constant_transformation. math-classes-8.19.0/implementations/000077500000000000000000000000001460576051100174125ustar00rootroot00000000000000math-classes-8.19.0/implementations/NType_naturals.v000066400000000000000000000135151460576051100225560ustar00rootroot00000000000000Require MathClasses.implementations.stdlib_binary_integers MathClasses.theory.integers MathClasses.orders.semirings. Require Import Coq.Setoids.Setoid Bignums.SpecViaZ.NSig Bignums.SpecViaZ.NSigNAxioms Coq.NArith.NArith Coq.ZArith.ZArith Coq.Program.Program Coq.Classes.Morphisms MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals MathClasses.interfaces.integers MathClasses.interfaces.orders MathClasses.interfaces.additional_operations. Module NType_Integers (Import anyN: NType). Module axioms := NTypeIsNAxioms anyN. #[global] Instance NType_equiv : Equiv t := eq. #[global] Instance NType_plus : Plus t := add. #[global] Instance NType_0 : Zero t := zero. #[global] Instance NType_1 : One t := one. #[global] Instance NType_mult : Mult t := mul. #[global] Instance: Setoid t | 10 := {}. #[global] Program Instance: ∀ x y: t, Decision (x = y) := λ x y, match compare x y with | Eq => left _ | _ => right _ end. Next Obligation. apply Zcompare_Eq_eq. now rewrite <-spec_compare. Qed. Next Obligation. rewrite spec_compare in *. intros E. apply Zcompare_Eq_iff_eq in E. auto. Qed. Ltac unfold_equiv := unfold equiv, NType_equiv, eq in *. Lemma NType_semiring_theory: semi_ring_theory zero one add mul eq. Proof. repeat split; repeat intro; axioms.zify; auto with zarith. Qed. #[global] Instance: SemiRing t | 10 := rings.from_stdlib_semiring_theory NType_semiring_theory. #[global] Instance inject_NType_N: Cast t N := to_N. #[global] Instance: Proper ((=) ==> (=)) to_N. Proof. intros x y E. unfold equiv, NType_equiv, eq in E. unfold to_N. now rewrite E. Qed. #[global] Instance: SemiRing_Morphism to_N. Proof. repeat (split; try apply _); unfold to_N; intros. now rewrite spec_add, Z2N.inj_add by apply spec_pos. unfold mon_unit, zero_is_mon_unit, NType_0. now rewrite spec_0. now rewrite spec_mul, Z2N.inj_mul by apply spec_pos. unfold mon_unit, one_is_mon_unit, NType_1. now rewrite spec_1. Qed. #[global] Instance inject_N_NType: Cast N t := of_N. #[global] Instance: Inverse to_N := of_N. #[global] Instance: Surjective to_N. Proof. split; try apply _. intros x y E. rewrite <-E. unfold to_N, inverse, compose. rewrite spec_of_N. apply N2Z.id. Qed. #[global] Instance: Injective to_N. Proof. split; try apply _. intros x y E. unfold equiv, NType_equiv, eq. unfold to_N in E. rewrite <-(Z2N.id (to_Z x)), <-(Z2N.id (to_Z y)) by now apply spec_pos. now rewrite E. Qed. #[global] Instance: Bijective to_N := {}. #[global] Instance: Inverse of_N := to_N. #[global] Instance: Bijective of_N. Proof. apply jections.flip_bijection. Qed. #[global] Instance: SemiRing_Morphism of_N. Proof. change (SemiRing_Morphism (to_N⁻¹)). split; apply _. Qed. #[global] Instance: NaturalsToSemiRing t := naturals.retract_is_nat_to_sr of_N. #[global] Instance: Naturals t := naturals.retract_is_nat of_N. #[global] Instance inject_NType_Z: Cast t Z := to_Z. #[global] Instance: Proper ((=) ==> (=)) to_Z. Proof. now intros x y E. Qed. #[global] Instance: SemiRing_Morphism to_Z. Proof. repeat (split; try apply _). exact spec_add. exact spec_0. exact spec_mul. exact spec_1. Qed. (* Order *) #[global] Instance NType_le: Le t := le. #[global] Instance NType_lt: Lt t := lt. #[global] Instance: Proper ((=) ==> (=) ==> iff) NType_le. Proof. intros ? ? E1 ? ? E2. unfold NType_le, le. unfold equiv, NType_equiv, eq in *. now rewrite E1, E2. Qed. #[global] Instance: SemiRingOrder NType_le. Proof. apply (semirings.projected_srorder to_Z). reflexivity. intros x y E. exists (sub y x). unfold_equiv. rewrite spec_add, spec_sub. rewrite Z.max_r by now apply Z.le_0_sub. ring. Qed. #[global] Instance: OrderEmbedding to_Z. Proof. now repeat (split; try apply _). Qed. #[global] Instance: TotalRelation NType_le. Proof. now apply (maps.projected_total_order to_Z). Qed. #[global] Instance: FullPseudoSemiRingOrder NType_le NType_lt. Proof. rapply semirings.dec_full_pseudo_srorder. intros x y. split. intro. split. apply axioms.lt_eq_cases. now left. intros E. destruct (irreflexivity (<) (to_Z x)). now rewrite E at 2. intros [E1 E2]. now destruct (proj1 (axioms.lt_eq_cases _ _) E1). Qed. (* Efficient comparison *) #[global] Program Instance: ∀ x y: t, Decision (x ≤ y) := λ x y, match (compare x y) with | Gt => right _ | _ => left _ end. Next Obligation. rewrite spec_compare in *. destruct (Z.compare_spec (to_Z x) (to_Z y)); try discriminate. now apply orders.lt_not_le_flip. Qed. Next Obligation. rewrite spec_compare in *. destruct (Z.compare_spec (to_Z x) (to_Z y)); try discriminate; try intuition. now apply Zeq_le. now apply orders.lt_le. Qed. Lemma NType_succ_1_plus x : succ x = 1 + x. Proof. unfold_equiv. rewrite spec_succ, rings.preserves_plus, rings.preserves_1. now rewrite commutativity. Qed. Lemma NType_two_2 : two = 2. Proof. unfold_equiv. rewrite spec_2. now rewrite rings.preserves_plus, rings.preserves_1. Qed. (* Efficient [nat_pow] *) #[global] Program Instance NType_pow: Pow t t := pow. #[global] Instance: NatPowSpec t t NType_pow. Proof. split. intros x1 y1 E1 x2 y2 E2. now apply axioms.pow_wd. intros x1. apply axioms.pow_0_r. intros x n. unfold_equiv. unfold "^", NType_pow. rewrite <-axioms.pow_succ_r by (red; rewrite spec_0; apply spec_pos). now rewrite NType_succ_1_plus. Qed. (* Efficient [shiftl] *) #[global] Program Instance NType_shiftl: ShiftL t t := shiftl. #[global] Instance: ShiftLSpec t t NType_shiftl. Proof. apply shiftl_spec_from_nat_pow. intros x y. unfold additional_operations.pow, NType_pow, additional_operations.shiftl, NType_shiftl. unfold_equiv. simpl. rewrite rings.preserves_mult, spec_pow. rewrite spec_shiftl, Z.shiftl_mul_pow2 by apply spec_pos. now rewrite <-NType_two_2, spec_2. Qed. (* Efficient [shiftr] *) #[global] Program Instance: ShiftR t t := shiftr. End NType_Integers. math-classes-8.19.0/implementations/QType_rationals.v000066400000000000000000000115741460576051100227270ustar00rootroot00000000000000Require MathClasses.theory.fields MathClasses.implementations.stdlib_rationals MathClasses.theory.int_pow. Require Import Coq.QArith.QArith Bignums.SpecViaQ.QSig MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.interfaces.integers MathClasses.interfaces.rationals MathClasses.interfaces.additional_operations MathClasses.theory.rings MathClasses.theory.rationals. Module QType_Rationals (Import anyQ: QType). Module Import props := QProperties anyQ. #[global] Instance QType_equiv: Equiv t := eq. #[global] Instance QType_plus: Plus t := add. #[global] Instance QType_0: Zero t := zero. #[global] Instance QType_1: One t := one. #[global] Instance QType_mult: Mult t := mul. #[global] Instance QType_negate: Negate t := opp. #[global] Instance QType_dec_recip: DecRecip t := inv. #[global] Instance: Setoid t := {}. #[global] Instance: ∀ x y: t, Decision (x = y). Proof with intuition. refine (λ x y, (match anyQ.eq_bool x y as p return p ≡ Qeq_bool (to_Q x) (to_Q y) → Decision (x = y) with | true => λ e, left _ | false => λ e, right _ end) (anyQ.spec_eq_bool x y)). (* hm, do we really need the anyQ.spec_eq_bool in here? *) apply Qeq_bool_iff... apply Qeq_bool_neq... Qed. (* We could get the above for free from the fact that anyQ.eq is just projected Qeq, but that mean that any comparison would involve two conversion to Q, which is a premature pessimization. *) Ltac unfold_equiv := unfold equiv, QType_equiv, eq. Add Ring Q : Qsrt. Lemma anyQ_field_theory: field_theory zero one add mul sub opp anyQ.div inv eq. (* No idea why this is missing in QSig. *) Proof. constructor. constructor; intros; qify; ring. exact neq_1_0. exact div_mul_inv. exact mul_inv_diag_l. Qed. #[global] Instance: DecField t. Proof. refine (dec_fields.from_stdlib_field_theory anyQ_field_theory _). unfold eq. now rewrite spec_inv, spec_0. Qed. (* Type-classified facts about to_Q/of_Q: *) #[global] Instance inject_QType_Q: Cast t Q := to_Q. #[global] Instance: Setoid_Morphism to_Q. Proof. constructor; try apply _. intros x y. auto. Qed. #[global] Instance: SemiRing_Morphism to_Q. Proof. repeat (constructor; try apply _); intros; qify; reflexivity. Qed. #[global] Instance inject_Q_QType: Cast Q t := of_Q. #[global] Instance: Inverse to_Q := of_Q. #[global] Instance: Surjective to_Q. Proof. constructor. intros x y E. rewrite <- E. apply spec_of_Q. apply _. Qed. #[global] Instance: Injective to_Q. Proof. constructor. auto. apply _. Qed. #[global] Instance: Bijective to_Q := {}. #[global] Instance: Inverse of_Q := to_Q. #[global] Instance: Bijective of_Q. Proof. apply jections.flip_bijection. Qed. #[global] Instance: SemiRing_Morphism of_Q. Proof. change (SemiRing_Morphism (to_Q⁻¹)). split; apply _. Qed. #[global] Instance: RationalsToFrac t := iso_to_frac of_Q. #[global] Instance: Rationals t := iso_is_rationals of_Q. (* Order *) #[global] Instance QType_le: Le t := le. #[global] Instance QType_lt: Lt t := lt. #[global] Instance: Proper ((=) ==> (=) ==> iff) QType_le. Proof. intros ? ? E1 ? ? E2. unfold QType_le, le, equiv, QType_equiv, eq in *. now rewrite E1, E2. Qed. #[global] Instance: SemiRingOrder QType_le. Proof. now apply (rings.projected_ring_order to_Q). Qed. #[global] Instance: OrderEmbedding to_Q. Proof. now repeat (split; try apply _). Qed. #[global] Instance: TotalRelation QType_le. Proof. now apply (maps.projected_total_order to_Q). Qed. #[global] Instance: FullPseudoSemiRingOrder QType_le QType_lt. Proof. rapply semirings.dec_full_pseudo_srorder. intros x y. change (to_Q x < to_Q y ↔ x ≤ y ∧ x ≠ y). now rewrite orders.lt_iff_le_ne. Qed. (* Efficient comparison *) #[global] Program Instance: ∀ x y: t, Decision (x ≤ y) := λ x y, match compare x y with | Gt => right _ | _ => left _ end. Next Obligation. rewrite spec_compare in *. destruct (Qcompare_spec (to_Q x) (to_Q y)); try discriminate. now apply orders.lt_not_le_flip. Qed. Next Obligation. rewrite spec_compare in *. destruct (Qcompare_spec (to_Q x) (to_Q y)); try discriminate; try intuition. now apply Zeq_le. now apply orders.lt_le. Qed. (* Efficient [int_pow] *) #[global] Program Instance QType_Zpow: Pow t Z := power. #[global] Instance: IntPowSpec t Z QType_Zpow. Proof. split; try apply _; unfold_equiv. intros. now rewrite spec_power, preserves_1. intros. rewrite spec_power, preserves_0. now apply int_pow_base_0. intros ? ? E. rewrite preserves_mult, 2!spec_power. rewrite preserves_0 in E. now apply int_pow_S. Qed. #[global] Program Instance QType_Npow: Pow t N := λ x n, power x (Z_of_N n). #[global] Instance: NatPowSpec t N QType_Npow. Proof. split; unfold "^", QType_Npow. solve_proper. intros. rewrite preserves_0. apply int_pow_0. intros ? ?. rewrite preserves_plus, preserves_1. apply int_pow.int_pow_S_nonneg. apply nat_int.to_semiring_nonneg. Qed. End QType_Rationals. math-classes-8.19.0/implementations/ZType_integers.v000066400000000000000000000163231460576051100225610ustar00rootroot00000000000000Require MathClasses.implementations.stdlib_binary_integers MathClasses.theory.integers MathClasses.orders.semirings. Require Import Bignums.SpecViaZ.ZSig Bignums.SpecViaZ.ZSigZAxioms Coq.NArith.NArith Coq.ZArith.ZArith MathClasses.implementations.nonneg_integers_naturals MathClasses.interfaces.orders MathClasses.interfaces.abstract_algebra MathClasses.interfaces.integers MathClasses.interfaces.additional_operations. Module ZType_Integers (Import anyZ: ZType). Module axioms := ZTypeIsZAxioms anyZ. #[global] Instance ZType_equiv : Equiv t := eq. #[global] Instance ZType_plus : Plus t := add. #[global] Instance ZType_0 : Zero t := zero. #[global] Instance ZType_1 : One t := one. #[global] Instance ZType_mult : Mult t := mul. #[global] Instance ZType_negate: Negate t := opp. #[global] Instance: Setoid t | 10 := {}. #[global] Program Instance: ∀ x y: t, Decision (x = y) := λ x y, match compare x y with | Eq => left _ | _ => right _ end. Next Obligation. apply Zcompare_Eq_eq. now rewrite <-spec_compare. Qed. Next Obligation. rewrite spec_compare in *. intros E. apply Zcompare_Eq_iff_eq in E. auto. Qed. Ltac unfold_equiv := unfold equiv, ZType_equiv, eq in *. Lemma ZType_ring_theory: ring_theory zero one add mul sub opp eq. Proof. repeat split; repeat intro; axioms.zify; auto with zarith. Qed. Local Instance: Ring t := rings.from_stdlib_ring_theory ZType_ring_theory. #[global] Instance inject_ZType_Z: Cast t Z := to_Z. #[global] Instance: Proper ((=) ==> (=)) to_Z. Proof. intros x y E. easy. Qed. #[global] Instance: SemiRing_Morphism to_Z. Proof. repeat (split; try apply _). exact spec_add. exact spec_0. exact spec_mul. exact spec_1. Qed. #[global] Instance inject_Z_ZType: Cast Z t := of_Z. #[global] Instance: Inverse to_Z := of_Z. #[global] Instance: Surjective to_Z. Proof. constructor. intros x y E. rewrite <- E. apply spec_of_Z. apply _. Qed. #[global] Instance: Injective to_Z. Proof. constructor. unfold_equiv. intuition. apply _. Qed. #[global] Instance: Bijective to_Z := {}. #[global] Instance: Inverse of_Z := to_Z. #[global] Instance: Bijective of_Z. Proof. apply jections.flip_bijection. Qed. #[global] Instance: SemiRing_Morphism of_Z. Proof. change (SemiRing_Morphism (to_Z⁻¹)). split; apply _. Qed. #[global] Instance: IntegersToRing t := integers.retract_is_int_to_ring of_Z. #[global] Instance: Integers t := integers.retract_is_int of_Z. (* Order *) #[global] Instance ZType_le: Le t := le. #[global] Instance ZType_lt: Lt t := lt. #[global] Instance: Proper ((=) ==> (=) ==> iff) ZType_le. Proof. intros ? ? E1 ? ? E2. unfold ZType_le, le. unfold equiv, ZType_equiv, eq in *. now rewrite E1, E2. Qed. #[global] Instance: SemiRingOrder ZType_le. Proof. now apply (rings.projected_ring_order to_Z). Qed. #[global] Instance: OrderEmbedding to_Z. Proof. now repeat (split; try apply _). Qed. #[global] Instance: TotalRelation ZType_le. Proof. now apply (maps.projected_total_order to_Z). Qed. #[global] Instance: FullPseudoSemiRingOrder ZType_le ZType_lt. Proof. rapply semirings.dec_full_pseudo_srorder. intros x y. change (to_Z x < to_Z y ↔ x ≤ y ∧ x ≠ y). now rewrite orders.lt_iff_le_ne. Qed. (* Efficient comparison *) #[global] Program Instance: ∀ x y: t, Decision (x ≤ y) := λ x y, match compare x y with | Gt => right _ | _ => left _ end. Next Obligation. rewrite spec_compare in *. destruct (Z.compare_spec (to_Z x) (to_Z y)); try discriminate. now apply orders.lt_not_le_flip. Qed. Next Obligation. rewrite spec_compare in *. destruct (Z.compare_spec (to_Z x) (to_Z y)); try discriminate; try intuition. now apply Zeq_le. now apply orders.lt_le. Qed. #[global] Program Instance: Abs t := abs. Next Obligation. split; intros E; unfold_equiv; rewrite spec_abs. apply Z.abs_eq. apply (order_preserving to_Z) in E. now rewrite rings.preserves_0 in E. rewrite rings.preserves_negate. apply Z.abs_neq. apply (order_preserving to_Z) in E. now rewrite rings.preserves_0 in E. Qed. (* Efficient division *) #[global] Instance ZType_div: DivEuclid t := anyZ.div. #[global] Instance ZType_mod: ModEuclid t := modulo. #[global] Instance: EuclidSpec t ZType_div ZType_mod. Proof. split; try apply _; unfold div_euclid, ZType_div. intros x y E. now apply axioms.div_mod. intros x y Ey. destruct (Z_mod_remainder (to_Z x) (to_Z y)) as [[Hl Hr] | [Hl Hr]]. intro. apply Ey. apply (injective to_Z). now rewrite rings.preserves_0. left; split. apply (order_reflecting to_Z). now rewrite spec_modulo, rings.preserves_0. apply (strictly_order_reflecting to_Z). now rewrite spec_modulo. right; split. apply (strictly_order_reflecting to_Z). now rewrite spec_modulo. apply (order_reflecting to_Z). now rewrite spec_modulo, rings.preserves_0. intros x. unfold_equiv. rewrite spec_div, rings.preserves_0. now apply Zdiv_0_r. intros x. unfold_equiv. unfold Nat.modulo. rewrite spec_modulo, rings.preserves_0. now apply Zmod_0_r. Qed. Lemma ZType_succ_1_plus x : succ x = 1 + x. Proof. unfold_equiv. rewrite spec_succ, rings.preserves_plus, rings.preserves_1. now rewrite commutativity. Qed. Lemma ZType_two_2 : two = 2. Proof. unfold_equiv. rewrite spec_2. now rewrite rings.preserves_plus, rings.preserves_1. Qed. (* Efficient [nat_pow] *) #[global] Program Instance ZType_pow: Pow t (t⁺) := pow. #[global] Instance: NatPowSpec t (t⁺) ZType_pow. Proof. split. intros x1 y1 E1 [x2] [y2] E2. now apply axioms.pow_wd. intros x1. apply axioms.pow_0_r. intros x [n ?]. unfold_equiv. unfold "^", ZType_pow. simpl. rewrite <-axioms.pow_succ_r; try easy. now rewrite ZType_succ_1_plus. Qed. #[global] Program Instance ZType_Npow: Pow t N := pow_N. #[global] Instance: NatPowSpec t N ZType_Npow. Proof. split; unfold "^", ZType_Npow. intros x1 y1 E1 x2 y2 E2. unfold_equiv. now rewrite 2!spec_pow_N, E1, E2. intros x1. unfold_equiv. now rewrite spec_pow_N, rings.preserves_1. intros x n. unfold_equiv. rewrite spec_mul, 2!spec_pow_N. rewrite rings.preserves_plus, Z.pow_add_r. now rewrite rings.preserves_1, Z.pow_1_r. easy. now apply Z_of_N_le_0. Qed. (* (* Efficient [log 2] *) Program Instance: Log (2:t) (t⁺) := log2. Next Obligation with auto. intros x. apply to_Z_Zle_sr_le. rewrite spec_log2, preserves_0. apply Z.log2_nonneg. Qed. Next Obligation with auto. intros [x Ex]. destruct (axioms.log2_spec x) as [E1 E2]. apply to_Z_sr_le_Zlt... unfold nat_pow, nat_pow_sig, ZType_pow; simpl. apply to_Z_Zle_sr_le in E1. apply to_Z_Zlt_sr_le in E2. rewrite ZType_two_2 in E1, E2. rewrite ZType_succ_plus_1, commutativity in E2... Qed. *) (* Efficient [shiftl] *) #[global] Program Instance ZType_shiftl: ShiftL t (t⁺) := shiftl. #[global] Instance: ShiftLSpec t (t⁺) ZType_shiftl. Proof. apply shiftl_spec_from_nat_pow. intros x [y Ey]. unfold additional_operations.pow, ZType_pow, additional_operations.shiftl, ZType_shiftl. unfold_equiv. simpl. rewrite rings.preserves_mult, spec_pow. rewrite spec_shiftl, Z.shiftl_mul_pow2. now rewrite <-ZType_two_2, spec_2. apply (order_preserving to_Z) in Ey. now rewrite rings.preserves_0 in Ey. Qed. (* Efficient [shiftr] *) #[global] Program Instance: ShiftR t (t⁺) := shiftr. End ZType_Integers. math-classes-8.19.0/implementations/bool.v000066400000000000000000000016441460576051100205410ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra. #[global] Instance bool_eq: Equiv bool := eq. #[global] Instance bool_bottom: Bottom bool := false. #[global] Instance bool_top: Top bool := true. #[global] Instance bool_join: Join bool := orb. #[global] Instance bool_meet: Meet bool := andb. #[global] Instance: BoundedJoinSemiLattice bool. Proof. repeat (split; try apply _); repeat intro. apply Bool.orb_assoc. apply Bool.orb_false_r. apply Bool.orb_comm. apply Bool.orb_diag. Qed. #[global] Instance: MeetSemiLattice bool. Proof. repeat (split; try apply _); repeat intro. apply Bool.andb_assoc. apply Bool.andb_comm. apply Bool.andb_diag. Qed. #[global] Instance: DistributiveLattice bool. Proof. repeat (split; try apply _); repeat intro. apply Bool.absoption_orb. apply Bool.absoption_andb. apply Bool.orb_andb_distrib_r. Qed. (* We don't have a boolean algebra class yet *) math-classes-8.19.0/implementations/dyadics.v000066400000000000000000000402121460576051100212200ustar00rootroot00000000000000(** The dyadic rationals are numbers of the shape [m * 2 ^ e] with [m : Z] and [e : Z] for some [Integers] implementation [Z]. These numbers form a ring and can be embedded into any [Rationals] implementation [Q]. *) Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.integers MathClasses.interfaces.naturals MathClasses.interfaces.rationals MathClasses.interfaces.additional_operations MathClasses.interfaces.orders MathClasses.orders.minmax MathClasses.orders.integers MathClasses.orders.rationals MathClasses.implementations.nonneg_integers_naturals MathClasses.implementations.stdlib_rationals MathClasses.theory.rationals MathClasses.theory.shiftl MathClasses.theory.int_pow MathClasses.theory.nat_pow MathClasses.theory.abs. Record Dyadic Z := dyadic { mant: Z; expo: Z }. Arguments dyadic {Z} _ _. Arguments mant {Z} _. Arguments expo {Z} _. Infix "▼" := dyadic (at level 80) : mc_scope. Section dyadics. Context `{Integers Z} `{Apart Z} `{!TrivialApart Z} `{!FullPseudoSemiRingOrder Zle Zlt} `{equiv_dec : ∀ (x y : Z), Decision (x = y)} `{le_dec : ∀ (x y : Z), Decision (x ≤ y)} `{!ShiftLSpec Z (Z⁺) sl}. Notation Dyadic := (Dyadic Z). Add Ring Z: (rings.stdlib_ring_theory Z). Global Program Instance dy_plus: Plus Dyadic := λ x y, if decide_rel (≤) (expo x) (expo y) then mant x + mant y ≪ (expo y - expo x)↾_ ▼ expo x ⊓ expo y else mant x ≪ (expo x - expo y)↾_ + mant y ▼ expo x ⊓ expo y. Next Obligation. now apply rings.flip_nonneg_minus. Qed. Next Obligation. apply rings.flip_nonneg_minus. now apply orders.le_flip. Qed. Global Instance dy_inject: Cast Z Dyadic := λ x, x ▼ 0. Global Instance dy_negate: Negate Dyadic := λ x, -mant x ▼ expo x. Global Instance dy_mult: Mult Dyadic := λ x y, mant x * mant y ▼ expo x + expo y. Global Instance dy_0: Zero Dyadic := ('0 : Dyadic). Global Instance dy_1: One Dyadic := ('1 : Dyadic). (* We define equality on the dyadics by injecting them into the rationals. Since we do not want to parametrize the equality relation and all properties on the dyadics by an arbitrary rationals implementations we have to pick some specific rationals implementations. Although [Frac Z] seems like a nice choice, it actually becomes quite slow, hence we stick with plain old [Qarith.Q]. Although we will make a specific choice for a rationals implementation to define the equality relation, we will define the embedding for arbitrary rationals implementations first. *) Section DtoQ_slow. Context `{Rationals Q} `{Pow Q Z} (ZtoQ: Z → Q). Definition DtoQ_slow (x : Dyadic) := ZtoQ (mant x) * 2 ^ (expo x). End DtoQ_slow. Section with_rationals. Context `{Rationals Q} `{!IntPowSpec Q Z ipw} `{!SemiRing_Morphism (ZtoQ: Z → Q)}. Add Ring Q: (rings.stdlib_ring_theory Q). Notation DtoQ_slow' := (DtoQ_slow ZtoQ). Lemma ZtoQ_shift (x n : Z) Pn : ZtoQ (x ≪ n↾Pn) = ZtoQ x * 2 ^ n. Proof. rewrite shiftl_nat_pow. rewrite rings.preserves_mult, nat_pow.preserves_nat_pow, rings.preserves_2. now rewrite <-(int_pow_nat_pow 2). Qed. Lemma DtoQ_slow_preserves_plus x y : DtoQ_slow' (x + y) = DtoQ_slow' x + DtoQ_slow' y. Proof. destruct x as [xn xe], y as [yn ye]. unfold plus at 1. unfold DtoQ_slow, dy_plus. simpl. case (decide_rel (≤) xe ye); intros E; simpl. rewrite rings.preserves_plus, ZtoQ_shift. rewrite (lattices.meet_l xe ye) by assumption. ring_simplify. rewrite <-associativity, <-int_pow_exp_plus. now setoid_replace (ye - xe + xe) with ye by ring. now apply (rings.is_ne_0 (2:Q)). rewrite rings.preserves_plus, ZtoQ_shift. rewrite lattices.meet_r. ring_simplify. rewrite <-associativity, <-int_pow_exp_plus. now setoid_replace (xe - ye + ye) with xe by ring. now apply (rings.is_ne_0 (2:Q)). now apply orders.le_flip. Qed. Lemma DtoQ_slow_preserves_negate x : DtoQ_slow' (-x) = -DtoQ_slow' x. Proof. unfold DtoQ_slow. simpl. rewrite rings.preserves_negate. ring. Qed. Lemma DtoQ_slow_preserves_mult x y : DtoQ_slow' (x * y) = DtoQ_slow' x * DtoQ_slow' y. Proof. destruct x as [xn xe], y as [yn ye]. unfold DtoQ_slow. simpl. rewrite rings.preserves_mult. rewrite int_pow_exp_plus. ring. apply (rings.is_ne_0 (2:Q)). Qed. Lemma DtoQ_slow_preserves_0 : DtoQ_slow' 0 = 0. Proof. unfold DtoQ_slow. simpl. rewrite rings.preserves_0. ring. Qed. Lemma DtoQ_slow_preserves_1 : DtoQ_slow' 1 = 1. Proof. unfold DtoQ_slow. simpl. rewrite int_pow_0, rings.preserves_1. ring. Qed. End with_rationals. Notation StdQ := QArith_base.Q. Notation ZtoStdQ := (integers.integers_to_ring Z StdQ). Notation DtoStdQ := (DtoQ_slow ZtoStdQ). Add Ring StdQ : (rings.stdlib_ring_theory StdQ). Global Instance dy_equiv: Equiv Dyadic := λ x y, DtoStdQ x = DtoStdQ y. Instance: Setoid Dyadic. Proof. now apply (setoids.projected_setoid DtoStdQ). Qed. Instance: Proper ((=) ==> (=)) DtoStdQ. Proof. now repeat red. Qed. Instance: Injective DtoStdQ. Proof. now repeat (split; try apply _). Qed. Global Instance: Ring Dyadic. Proof. apply (rings.projected_ring DtoStdQ). exact DtoQ_slow_preserves_plus. exact DtoQ_slow_preserves_0. exact DtoQ_slow_preserves_mult. exact DtoQ_slow_preserves_1. exact DtoQ_slow_preserves_negate. Qed. Global Instance dyadic_proper: Proper ((=) ==> (=) ==> (=)) dyadic. Proof. intros ? ? E1 ? ? E2. unfold equiv, dy_equiv, DtoQ_slow. simpl. now rewrite E1, E2. Qed. Instance: SemiRing_Morphism DtoStdQ. Proof. repeat (split; try apply _). exact DtoQ_slow_preserves_plus. exact DtoQ_slow_preserves_0. exact DtoQ_slow_preserves_mult. exact DtoQ_slow_preserves_1. Qed. Instance: Setoid_Morphism dy_inject. Proof. split; try apply _. intros x y E. unfold equiv, dy_equiv, dy_inject, DtoQ_slow. simpl in *. rewrite int_pow_0. now rewrite E. Qed. Global Instance: Injective dy_inject. Proof. repeat (split; try apply _). intros x y E. unfold equiv, dy_equiv, dy_inject, DtoQ_slow in E. simpl in *. rewrite int_pow_0 in E. ring_simplify in E. now apply (injective ZtoStdQ). Qed. Global Instance: SemiRing_Morphism dy_inject. Proof. repeat (split; try apply _). intros x y. unfold sg_op at 2, plus_is_sg_op, dy_plus. unfold equiv, dy_equiv, dy_inject, DtoQ_slow; simpl. case (decide_rel); intros E; simpl. rewrite 2!rings.preserves_plus, ZtoQ_shift. rewrite rings.plus_negate_r. rewrite lattices.meet_l, int_pow_0. ring. reflexivity. now destruct E. intros x y. unfold sg_op at 2, mult_is_sg_op, dy_mult. simpl. now setoid_replace (0 + 0:Z) with (0:Z) by ring. Qed. Lemma dy_eq_dec_aux (x y : Dyadic) p : mant x = mant y ≪ (expo y - expo x)↾p ↔ x = y. Proof. destruct x as [xm xe], y as [ym ye]. assert (xe ≤ ye). now apply rings.flip_nonneg_minus. split; intros E. unfold equiv, dy_equiv, DtoQ_slow. simpl in *. rewrite E, ZtoQ_shift. rewrite <-associativity, <-int_pow_exp_plus. now setoid_replace (ye - xe + xe) with ye by ring. easy. unfold equiv, dy_equiv, DtoQ_slow in E. simpl in *. apply (injective ZtoStdQ). apply (right_cancellation (.*.) (2 ^ xe)). rewrite E, ZtoQ_shift. rewrite <-associativity, <-int_pow_exp_plus. now setoid_replace (ye - xe + xe) with ye by ring. easy. Qed. Lemma dy_eq_dec_aux_neg (x y : Dyadic) p : mant x ≠ mant y ≪ (expo y - expo x)↾p ↔ x ≠ y. Proof. split; intros E; intro; apply E; eapply dy_eq_dec_aux; eassumption. Qed. Global Program Instance dy_eq_dec : ∀ (x y: Dyadic), Decision (x = y) := λ x y, if decide_rel (≤) (expo x) (expo y) then if decide_rel (=) (mant x) (mant y ≪ (expo y - expo x)↾_) then left _ else right _ else if decide_rel (=) (mant x ≪ (expo x - expo y)↾_) (mant y) then left _ else right _. Next Obligation. now apply rings.flip_nonneg_minus. Qed. Next Obligation. eapply dy_eq_dec_aux; eauto. Qed. Next Obligation. eapply dy_eq_dec_aux_neg; eauto. Qed. Next Obligation. apply rings.flip_nonneg_minus. now apply orders.le_flip. Qed. Next Obligation. symmetry. eapply dy_eq_dec_aux. symmetry. eassumption. Qed. Next Obligation. apply not_symmetry. eapply dy_eq_dec_aux_neg. apply not_symmetry. eassumption. Qed. Global Instance dy_pow `{!Pow Z (Z⁺)} : Pow Dyadic (Z⁺) := λ x n, (mant x) ^ n ▼ 'n * expo x. Global Instance dy_pow_spec `{!NatPowSpec Z (Z⁺) pw} : NatPowSpec Dyadic (Z⁺) dy_pow. Proof. split; unfold pow, dy_pow. intros [xm xe] [ym ye] E1 e1 e2 E2. unfold equiv, dy_equiv, DtoQ_slow in E1 |- *. simpl in *. setoid_replace (xm ^ e1) with (xm ^ e2) by now apply (_ : Proper ((=) ==> (=) ==> (=)) pw). (* fixme *) rewrite E2. clear e1 E2. rewrite 2!(preserves_nat_pow (f:=ZtoStdQ)). rewrite 2!(commutativity ('e2 : Z)). rewrite 2!int_pow_exp_mult. rewrite 2!int_pow_nat_pow. rewrite <-2!nat_pow_base_mult. now rewrite E1. intros [xm xe]. simpl. rewrite rings.preserves_0, left_absorb. now rewrite nat_pow_0. intros [xm xe] n. simpl. rewrite nat_pow_S. rewrite rings.preserves_plus, rings.preserves_1. now rewrite distribute_r, left_identity. Qed. Global Instance dy_shiftl: ShiftL Dyadic Z := λ x n, mant x ▼ n + expo x. Global Instance: ShiftLSpec Dyadic Z dy_shiftl. Proof. split. intros [xm xe] [ym ye] E1 n1 n2 E2. unfold shiftl, dy_shiftl, equiv, dy_equiv, DtoQ_slow in *. simpl in *. rewrite 2!int_pow_exp_plus; try apply (rings.is_ne_0 (2:StdQ)). transitivity (ZtoStdQ xm * 2 ^ xe * 2 ^ n1). ring. rewrite E1, E2. ring. intros [xm xe]. unfold shiftl, dy_shiftl, equiv, dy_equiv, DtoQ_slow. simpl. now rewrite left_identity. intros [xm xe] n. simpl. rewrite <-(rings.preserves_2 (f:=dy_inject)). unfold shiftl, dy_shiftl, equiv, dy_equiv, DtoQ_slow. simpl. rewrite <-associativity, int_pow_S. rewrite rings.preserves_mult, rings.preserves_2. rewrite rings.plus_0_l. ring. apply (rings.is_ne_0 (2:StdQ)). Qed. Global Instance dy_le: Le Dyadic := λ x y, DtoStdQ x ≤ DtoStdQ y. Global Instance dy_lt: Lt Dyadic := orders.dec_lt. Instance: Proper ((=) ==> (=) ==> iff) dy_le. Proof. intros [x1m x1e] [y1m y1e] E1 [x2m x2e] [y2m y2e] E2. unfold dy_le, equiv, dy_equiv, DtoQ_slow in *. simpl in *. now rewrite E1, E2. Qed. Instance: SemiRingOrder dy_le. Proof. now apply (rings.projected_ring_order DtoStdQ). Qed. Instance: TotalRelation dy_le. Proof. now apply (maps.projected_total_order DtoStdQ). Qed. Instance: OrderEmbedding DtoStdQ. Proof. now repeat (split; try apply _). Qed. Global Instance: ZeroProduct Dyadic. Proof. intros x y E. destruct (zero_product (DtoStdQ x) (DtoStdQ y)) as [Ex|Ey]. now rewrite <-rings.preserves_mult, E, rings.preserves_0. left. apply (injective DtoStdQ). now rewrite Ex, rings.preserves_0. right. apply (injective DtoStdQ). now rewrite Ey, rings.preserves_0. Qed. Global Instance: FullPseudoSemiRingOrder dy_le dy_lt. Proof. now rapply (semirings.dec_full_pseudo_srorder (R:=Dyadic)). Qed. Lemma nonneg_mant (x : Dyadic) : 0 ≤ x ↔ 0 ≤ mant x. Proof. split; intros E. unfold le, dy_le, DtoQ_slow in E. simpl in *. apply (order_reflecting ZtoStdQ). apply (order_reflecting (.* 2 ^ expo x)). now rewrite rings.preserves_0, left_absorb in E |- *. unfold le, dy_le, DtoQ_slow. simpl. apply (order_preserving ZtoStdQ) in E. apply (order_preserving (.* 2 ^ expo x)) in E. now rewrite rings.preserves_0, left_absorb in E |- *. Qed. Lemma nonpos_mant (x : Dyadic) : x ≤ 0 ↔ mant x ≤ 0. Proof. rewrite 2!rings.flip_nonpos_negate. apply nonneg_mant. Qed. Global Program Instance dy_abs `{!Abs Z} : Abs Dyadic := λ x, abs (mant x) ▼ expo x. Next Obligation. split; intros E. rewrite abs_nonneg. now destruct x. now apply nonneg_mant. rewrite abs_nonpos. now destruct x. now apply nonpos_mant. Qed. Lemma dy_le_dec_aux (x y : Dyadic) p : mant x ≤ mant y ≪ (expo y - expo x)↾p → x ≤ y. Proof. destruct x as [xm xe], y as [ym ye]. intros E. unfold le, dy_le, DtoQ_slow. simpl in *. apply (order_preserving ZtoStdQ) in E. rewrite ZtoQ_shift in E. apply (order_preserving (.* 2 ^ xe)) in E. rewrite <-associativity, <-int_pow_exp_plus in E. now setoid_replace ((ye - xe) + xe) with ye in E by ring. now apply (rings.is_ne_0 (2:StdQ)). Qed. Local Obligation Tactic := idtac. Global Program Instance dy_le_dec : ∀ (x y: Dyadic), Decision (x ≤ y) := λ x y, if decide_rel (≤) (expo x) (expo y) then if decide_rel (≤) (mant x) (mant y ≪ (expo y - expo x)↾_) then left _ else right _ else if decide_rel (≤) (mant x ≪ (expo x - expo y)↾_) (mant y) then left _ else right _. Next Obligation. intros. now apply rings.flip_nonneg_minus. Qed. Next Obligation. intros x y E1 E2. eapply dy_le_dec_aux. eassumption. Qed. Next Obligation. intros x y E1 E2. apply orders.lt_not_le_flip. apply orders.not_le_lt_flip in E2. apply rings.flip_lt_negate in E2. rewrite orders.lt_iff_le_ne in E2. destruct E2 as [E2a E2b]. split. apply rings.flip_le_negate. eapply dy_le_dec_aux. simpl. rewrite shiftl_negate. eassumption. intros E3. apply E2b. apply sm_proper. apply dy_eq_dec_aux. now symmetry. Qed. Next Obligation. intros. apply rings.flip_nonneg_minus. now apply orders.le_flip. Qed. Next Obligation. intros x y E1 E2. apply orders.le_equiv_lt in E2. destruct E2 as [E2 | E2]. apply orders.eq_le. symmetry in E2 |- *. eapply dy_eq_dec_aux. eassumption. apply rings.flip_le_negate. eapply dy_le_dec_aux. simpl. rewrite shiftl_negate. apply rings.flip_le_negate. apply orders.lt_le, E2. Qed. Next Obligation. intros x y E1 E2. apply orders.not_le_lt_flip in E2. rewrite orders.lt_iff_le_ne in E2. destruct E2 as [E2a E2b]. apply orders.lt_not_le_flip. apply orders.lt_iff_le_ne. split. eapply dy_le_dec_aux. eassumption. eapply dy_eq_dec_aux_neg. eassumption. Qed. (* The embedding into the rationals as presented above, is quite slow because it involves performing exponentiation on the rationals while we have a (presumably) much faster shift operation on the integers. We therefore define a faster embedding, prove that it is equivalent to the one above, and derive some basic properties. *) Section DtoQ. Context `{Rationals Q} (ZtoQ: Z → Q) `{!SemiRing_Morphism (ZtoQ: Z → Q)}. Local Obligation Tactic := program_simpl. Program Definition DtoQ (x : Dyadic) : Q := if decide_rel (≤) 0 (expo x) then ZtoQ (mant x ≪ (expo x)↾_) else ZtoQ (mant x) / ZtoQ (1 ≪ (-expo x)↾_). Next Obligation. apply rings.flip_nonpos_negate. now apply orders.le_flip. Qed. End DtoQ. Section embed_rationals. Context `{Rationals Q} `{!IntPowSpec Q Z ipw} `{!SemiRing_Morphism (ZtoQ: Z → Q)}. Context `{Apart Q} `{!TrivialApart Q} `{!FullPseudoSemiRingOrder Qlt Qle}. Add Ring Q2 : (rings.stdlib_ring_theory Q). Notation DtoQ' := (DtoQ ZtoQ). Notation DtoQ_slow' := (DtoQ_slow ZtoQ). Notation StdQtoQ := (rationals_to_rationals StdQ Q). Instance: Params (@DtoQ_slow) 6 := {}. Lemma DtoQ_slow_correct : DtoQ_slow' = StdQtoQ ∘ DtoStdQ. Proof. intros x y E. unfold compose. rewrite <- E. clear y E. unfold DtoQ_slow. rewrite rings.preserves_mult, (preserves_int_pow 2), rings.preserves_2. now rewrite (integers.to_ring_unique_alt ZtoQ (StdQtoQ ∘ ZtoStdQ)). Qed. Global Instance: SemiRing_Morphism DtoQ_slow'. Proof. apply (rings.semiring_morphism_proper _ _ DtoQ_slow_correct), _. Qed. Global Instance: Injective DtoQ_slow'. Proof. rewrite DtoQ_slow_correct. apply _. Qed. Global Instance: OrderEmbedding DtoQ_slow'. Proof. apply (maps.order_embedding_proper _ _ DtoQ_slow_correct). apply _. Qed. Lemma DtoQ_correct : DtoQ' = DtoQ_slow'. Proof. intros x y E. rewrite <-E. clear y E. unfold DtoQ, DtoQ_slow. destruct x as [xm xe]. simpl. case (decide_rel); intros E. now rewrite ZtoQ_shift. rewrite int_pow_negate_alt, ZtoQ_shift. now rewrite rings.preserves_1, left_identity. Qed. Global Instance: SemiRing_Morphism DtoQ'. Proof. apply (rings.semiring_morphism_proper _ _ DtoQ_correct), _. Qed. Global Instance: Injective DtoQ'. Proof. rewrite DtoQ_correct. apply _. Qed. Global Instance: OrderEmbedding DtoQ'. Proof. apply (maps.order_embedding_proper _ _ DtoQ_correct). apply _. Qed. End embed_rationals. End dyadics. math-classes-8.19.0/implementations/fast_integers.v000066400000000000000000000022161460576051100224370ustar00rootroot00000000000000Require Import Bignums.BigZ.BigZ MathClasses.interfaces.abstract_algebra MathClasses.interfaces.integers MathClasses.interfaces.additional_operations MathClasses.implementations.fast_naturals. Require Export MathClasses.implementations.ZType_integers. Module BigZ_Integers := ZType_Integers BigZ. #[global] Instance inject_fastN_fastZ: Cast bigN bigZ := BigZ.Pos. #[global] Instance: SemiRing_Morphism inject_fastN_fastZ. Proof. repeat (split; try apply _); intuition. Qed. #[global] Program Instance bigZ_pow: Pow bigZ bigN := λ x n, BigZ.pow x ('n). #[global] Instance: NatPowSpec bigZ bigN _. Proof. split; unfold pow, bigZ_pow. solve_proper. intro. apply BigZ.pow_0_r. intros x n. rewrite rings.preserves_plus, rings.preserves_1, BigZ.add_1_l. apply BigZ.pow_succ_r. change (0 ≤ cast bigN bigZ n). now apply nat_int.to_semiring_nonneg. Qed. #[global] Instance fastZ_shiftl: ShiftL bigZ bigN := λ x n, BigZ.shiftl x ('n). #[global] Instance: ShiftLSpec bigZ bigN _. Proof. apply shiftl_spec_from_nat_pow. intros. apply BigZ.shiftl_mul_pow2. change (0 ≤ cast bigN bigZ n). now apply nat_int.to_semiring_nonneg. Qed. math-classes-8.19.0/implementations/fast_naturals.v000066400000000000000000000002571460576051100224530ustar00rootroot00000000000000Require Import Bignums.BigN.BigN MathClasses.interfaces.naturals. Require Export MathClasses.implementations.NType_naturals. Module BigN_Integers := NType_Integers BigN. math-classes-8.19.0/implementations/fast_rationals.v000066400000000000000000000127021460576051100226140ustar00rootroot00000000000000Require MathClasses.theory.rings MathClasses.theory.shiftl MathClasses.theory.int_pow. Require Import Coq.QArith.QArith Bignums.BigQ.BigQ MathClasses.interfaces.orders MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals MathClasses.interfaces.integers MathClasses.interfaces.rationals MathClasses.interfaces.additional_operations MathClasses.implementations.fast_naturals MathClasses.implementations.fast_integers MathClasses.implementations.field_of_fractions MathClasses.implementations.stdlib_rationals. Require Export MathClasses.implementations.QType_rationals. Module Import BigQ_Rationals := QType_Rationals BigQ. (* Embedding of [bigZ] into [bigQ] *) #[global] Instance inject_bigZ_bigQ: Cast bigZ bigQ := BigQ.Qz. #[global] Instance inject_bigN_bigQ: Cast bigN bigQ := cast bigZ bigQ ∘ cast bigN bigZ. #[global] Instance inject_Z_bigQ: Cast Z bigQ := cast bigZ bigQ ∘ cast Z bigZ. #[global] Instance: Proper ((=) ==> (=)) inject_bigZ_bigQ. Proof. intros x y E. unfold_equiv. unfold Qeq. simpl. now rewrite E. Qed. #[global] Instance: SemiRing_Morphism inject_bigZ_bigQ. Proof. repeat (split; try apply _). Qed. #[global] Instance: SemiRing_Morphism inject_bigN_bigQ. Proof. unfold inject_bigN_bigQ. apply _. Qed. #[global] Instance: SemiRing_Morphism inject_Z_bigQ. Proof. unfold inject_Z_bigQ. apply _. Qed. #[global] Instance: Proper ((=) ==> (=) ==> (=)) BigQ.Qq. Proof. intros x1 y1 E1 x2 y2 E2. do 4 red. simpl. case_eq (BigN.eqb x2 BigN.zero); intros Ex2; case_eq (BigN.eqb y2 BigN.zero); intros Ey2. reflexivity. rewrite E2 in Ex2. edestruct eq_true_false_abs; eassumption. rewrite E2 in Ex2. edestruct eq_true_false_abs; eassumption. simpl. do 3 red in E1. do 3 red in E2. now rewrite E1, E2. Qed. (* Why is the below not in the stdlib? *) Lemma bigQ_div_bigQq (n : bigZ) (d : bigN) : BigQ.Qq n d = 'n / 'd. Proof. change (n # d == 'n / (BigQ.Qz (BigZ.Pos d)))%bigQ. unfold BigQ.div, BigQ.inv. case_eq (BigZ.zero ?= BigZ.Pos d)%bigZ; intros Ed. transitivity BigQ.zero; [| ring]. do 2 red. simpl. case_eq (BigN.eqb d BigN.zero); intros Ed2; [reflexivity |]. rewrite BigZ.spec_compare in Ed. destruct (proj2 (not_true_iff_false _) Ed2). apply BigN.eqb_eq. symmetry. now apply Zcompare_Eq_eq. unfold BigQ.mul. simpl. rewrite right_identity. reflexivity. destruct (BigZ.compare_spec BigZ.zero (BigZ.Pos d)); try discriminate. destruct (orders.lt_not_le_flip 0 ('d : bigZ)); trivial. now apply nat_int.to_semiring_nonneg. Qed. Lemma bigQ_div_bigQq_alt (n : bigZ) (d : bigN) : BigQ.Qq n d = 'n / 'cast bigN bigZ d. Proof. apply bigQ_div_bigQq. Qed. (* Embedding of [bigQ] into [Frac bigZ] *) #[global] Instance inject_bigQ_frac_bigZ: Cast bigQ (Frac bigZ) := λ x, match x with | BigQ.Qz n => 'n | BigQ.Qq n d => match decide_rel (=) ('d : bigZ) 0 with | left _ => 0 | right E => frac n ('d) E end end. Lemma inject_bigQ_frac_bigZ_correct: cast bigQ (Frac bigZ) = rationals_to_frac bigQ bigZ. Proof. intros x y E. rewrite <-E. clear y E. destruct x as [n|n d]. rapply (integers.to_ring_unique_alt (cast bigZ (Frac bigZ)) (rationals_to_frac bigQ bigZ ∘ cast bigZ bigQ)). unfold cast at 1. simpl. rewrite bigQ_div_bigQq_alt. rewrite rings.preserves_mult, dec_fields.preserves_dec_recip. case (decide_rel (=) ('d : bigZ) 0); intros Ed. rewrite Ed. rewrite !rings.preserves_0. rewrite dec_recip_0. now rewrite rings.mult_0_r. rewrite 2!(integers.to_ring_twice _ _ (cast bigZ (Frac bigZ))). now rewrite Frac_dec_mult_num_den at 1. Qed. #[global] Instance: Injective inject_bigQ_frac_bigZ. Proof. rewrite inject_bigQ_frac_bigZ_correct. apply _. Qed. #[global] Instance: SemiRing_Morphism inject_bigQ_frac_bigZ. Proof. eapply rings.semiring_morphism_proper. apply inject_bigQ_frac_bigZ_correct. apply _. Qed. (* Efficient shiftl on [bigQ] *) #[global] Instance bigQ_shiftl: ShiftL bigQ bigZ := λ x k, match k with | BigZ.Pos k => match x with | BigQ.Qz n => '(n ≪ k) | BigQ.Qq n d => BigQ.Qq (n ≪ k) d end | BigZ.Neg k => match x with | BigQ.Qz n => BigQ.Qq n (1 ≪ k) | BigQ.Qq n d => BigQ.Qq n (d ≪ k) end end. #[global] Instance: ShiftLSpec bigQ bigZ _. Proof. apply shiftl_spec_from_int_pow. unfold shiftl, bigQ_shiftl. intros [n|n d] [k|k]. rewrite shiftl.preserves_shiftl. now rewrite <-shiftl.shiftl_int_pow, shiftl.preserves_shiftl_exp. change (BigZ.Neg k) with ((-'k)%mc:bigZ). rewrite int_pow.int_pow_negate. rewrite bigQ_div_bigQq, shiftl.preserves_shiftl. rewrite <-(shiftl.preserves_shiftl_exp (f:=cast bigN bigZ)). now rewrite rings.preserves_1, shiftl.shiftl_base_1_int_pow. rewrite 2!bigQ_div_bigQq. rewrite shiftl.preserves_shiftl. rewrite <-(shiftl.preserves_shiftl_exp (f:=cast bigN bigZ)). rewrite shiftl.shiftl_int_pow. now rewrite <-2!associativity, (commutativity (/cast bigN bigQ d)). change (BigZ.Neg k) with ((-'k)%mc : bigZ). rewrite int_pow.int_pow_negate. rewrite 2!bigQ_div_bigQq. rewrite shiftl.preserves_shiftl. rewrite <-(shiftl.preserves_shiftl_exp (f:=cast bigN bigZ)). rewrite shiftl.shiftl_int_pow. now rewrite dec_fields.dec_recip_distr, associativity. Qed. #[global] Instance bigQ_Zshiftl: ShiftL bigQ Z := λ x k, x ≪ 'k. #[global] Instance: ShiftLSpec bigQ Z _. Proof. split; unfold shiftl, bigQ_Zshiftl. solve_proper. intro. now apply shiftl_0. intros x n. rewrite rings.preserves_plus. now apply shiftl_S. Qed. math-classes-8.19.0/implementations/field_of_fractions.v000066400000000000000000000136201460576051100234220ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.theory.rings MathClasses.theory.dec_fields. Inductive Frac R `{Rap : Equiv R} `{Rzero : Zero R} : Type := frac { num: R; den: R; den_ne_0: den ≠ 0 }. (* We used to have [den] and [den_nonzero] bundled, which did work relatively nicely with Program, but the extra messyness in proofs etc turned out not to be worth it. *) Arguments frac {R Rap Rzero} _ _ _. Arguments num {R Rap Rzero} _. Arguments den {R Rap Rzero} _. Arguments den_ne_0 {R Rap Rzero} _ _. Section contents. Context `{IntegralDomain R} `{∀ x y, Decision (x = y)}. Add Ring R: (stdlib_ring_theory R). Global Instance Frac_equiv : Equiv (Frac R) | 0 := λ x y, num x * den y = num y * den x. Instance: Setoid (Frac R). Proof with auto. split; red; unfold equiv, Frac_equiv. reflexivity. intros x y E. now symmetry. intros [nx dx] [ny dy] [nz dz] V W. simpl in *. apply (left_cancellation_ne_0 (.*.) dy)... rewrite 2!associativity. rewrite 2!(commutativity dy). rewrite V, <- W. ring. Qed. Global Instance Frac_dec : ∀ x y: Frac R, Decision (x = y) := λ x y, decide_rel (=) (num x * den y) (num y * den x). (* injection from R *) Global Program Instance Frac_inject: Cast R (Frac R) := λ r, frac r 1 _. Next Obligation. exact (is_ne_0 1). Qed. Instance: Proper ((=) ==> (=)) Frac_inject. Proof. intros x1 x2 E. unfold equiv, Frac_equiv. simpl. now rewrite E. Qed. (* Relations, operations and constants *) Global Program Instance Frac_plus: Plus (Frac R) := λ x y, frac (num x * den y + num y * den x) (den x * den y) _. Next Obligation. destruct x, y. simpl. now apply mult_ne_0. Qed. Global Instance Frac_0: Zero (Frac R) := ('0 : Frac R). Global Instance Frac_1: One (Frac R) := ('1 : Frac R). Global Instance Frac_negate: Negate (Frac R) := λ x, frac (- num x) (den x) (den_ne_0 x). Global Program Instance Frac_mult: Mult (Frac R) := λ x y, frac (num x * num y) (den x * den y) _. Next Obligation. destruct x, y. simpl. now apply mult_ne_0. Qed. Ltac unfolds := unfold Frac_negate, Frac_plus, equiv, Frac_equiv in *; simpl in *. Ltac ring_on_ring := repeat intro; unfolds; try ring. Lemma Frac_nonzero_num x : x ≠ 0 ↔ num x ≠ 0. Proof. split; intros E F; apply E; unfolds. rewrite F. ring. rewrite right_identity in F. rewrite F. apply left_absorb. Qed. Instance: Proper ((=) ==> (=) ==> (=)) Frac_plus. Proof with try ring. intros x x' E y y' E'. unfolds. transitivity (num x * den x' * den y * den y' + num y * den y' * den x * den x')... rewrite E, E'... Qed. Instance: Proper ((=) ==> (=)) Frac_negate. Proof. intros x y E. unfolds. rewrite <-negate_mult_distr_l, E. ring. Qed. Instance: Proper ((=) ==> (=) ==> (=)) Frac_mult. Proof with try ring. intros x y E x0 y0 E'. unfolds. transitivity (num x * den y * (num x0 * den y0))... rewrite E, E'... Qed. Global Instance: Ring (Frac R). Proof. repeat (split; try apply _); ring_on_ring. Qed. Global Instance Frac_dec_recip: DecRecip (Frac R) := λ x, match decide_rel (=) (num x) 0 with | left _ => 0 | right P => frac (den x) (num x) P end. Instance: Setoid_Morphism Frac_dec_recip. Proof. split; try apply _. intros [xn xd Px] [yn yd Py]. unfolds. unfold Frac_dec_recip. simpl. case (decide_rel (=) xn 0); case (decide_rel (=) yn 0); intros Ey Ex; simpl. reflexivity. rewrite Ex. intros E. destruct Ey. apply (right_cancellation_ne_0 (.*.) xd); trivial. rewrite <-E. ring. rewrite Ey. intros E. destruct Ex. apply (right_cancellation_ne_0 (.*.) yd); trivial. rewrite E. ring. symmetry. now rewrite (commutativity yd), (commutativity xd). Qed. Global Instance: DecField (Frac R). Proof. constructor; try apply _. red. unfolds. rewrite 2!mult_1_r. apply (is_ne_0 1). unfold dec_recip, Frac_dec_recip. case (decide_rel); simpl; intuition. intros [xn xs] Ex. unfold dec_recip, Frac_dec_recip. simpl. case (decide_rel); simpl. intros E. destruct Ex. unfolds. rewrite E. ring. intros. ring_on_ring. Qed. Lemma Frac_dec_mult_num_den x : x = 'num x / 'den x. Proof. unfold dec_recip, Frac_dec_recip. simpl. case (decide_rel); simpl; intros E. now destruct (den_ne_0 x). unfolds. ring. Qed. (* A final word about inject *) Global Instance: SemiRing_Morphism Frac_inject. Proof. repeat (constructor; try apply _); try reflexivity. intros x y. change ((x + y) * (1 * 1) = (x * 1 + y * 1) * 1). ring. intros x y. change ((x * y) * (1 * 1) = x * y * 1). ring. Qed. Global Instance: Injective Frac_inject. Proof. repeat (constructor; try apply _). intros x y. unfolds. rewrite 2!mult_1_r. intuition. Qed. End contents. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Opaque Frac_equiv. Set Warnings "+unsupported-attributes". Section morphisms. Context `{IntegralDomain R1} `{∀ x y : R1, Decision (x = y)}. Context `{IntegralDomain R2} `{∀ x y : R2, Decision (x = y)}. Context `(f : R1 → R2) `{!SemiRing_Morphism f} `{!Injective f}. Program Definition Frac_lift (x : Frac R1) : Frac R2 := frac (f (num x)) (f (den x)) _. Next Obligation. apply injective_ne_0. now apply (den_ne_0 x). Qed. Instance: Proper ((=) ==> (=)) Frac_lift. Proof. intros x y E. unfold equiv, Frac_equiv, Frac_lift in *. simpl. now rewrite <-2!preserves_mult, E. Qed. Global Instance: SemiRing_Morphism Frac_lift. Proof. pose proof (_:Ring (Frac R1)). repeat (split; try apply _); unfold equiv, Frac_equiv, Frac_lift in *; simpl. intros x y. now rewrite preserves_plus, ?preserves_mult. now rewrite preserves_0, preserves_1. intros x y. now rewrite ?preserves_mult. now rewrite preserves_1. Qed. Global Instance: Injective Frac_lift. Proof. split; try apply _. intros x y E. unfold equiv, Frac_equiv, Frac_lift in *. simpl in *. apply (injective f). now rewrite 2!preserves_mult. Qed. End morphisms. math-classes-8.19.0/implementations/intfrac_rationals.v000066400000000000000000000006721460576051100233100ustar00rootroot00000000000000Require Import MathClasses.interfaces.rationals MathClasses.interfaces.integers MathClasses.interfaces.abstract_algebra MathClasses.theory.rationals. Require Export MathClasses.implementations.field_of_fractions. Section intfrac_rationals. Context `{Integers Z}. Global Instance: RationalsToFrac (Frac Z) := alt_to_frac id. Global Instance: Rationals (Frac Z) := alt_Build_Rationals id (cast Z (Frac Z)). End intfrac_rationals. math-classes-8.19.0/implementations/list.v000066400000000000000000000270461460576051100205650ustar00rootroot00000000000000Require Import Coq.Lists.List Coq.Lists.SetoidList MathClasses.interfaces.abstract_algebra MathClasses.interfaces.monads MathClasses.theory.monads. Import ListNotations. (* The below belongs in the stdlib *) Section equivlistA_misc. Context `{Equivalence A eqA}. Lemma NoDupA_singleton x : NoDupA eqA [x]. Proof. apply NoDupA_cons. intros E. inversion E. auto. Qed. Global Instance equivlistA_cons_proper: Proper (eqA ==> equivlistA eqA ==> equivlistA eqA) cons. Proof. assert (∀ x₁ x₂ l₁ l₂ x, eqA x₁ x₂ → equivlistA eqA l₁ l₂ → InA eqA x (x₁ :: l₁) → InA eqA x (x₂ :: l₂)) as aux. intros ? ? ? ? ? E1 E2 E3. inversion_clear E3. apply InA_cons_hd. now rewrite <-E1. apply InA_cons_tl. now rewrite <-E2. split; eapply aux; auto; now symmetry. Qed. Lemma InA_singleton x y : InA eqA x [y] ↔ eqA x y. Proof. split; intros E. inversion_clear E; auto. now destruct (proj1 (InA_nil eqA x)). rewrite E. intuition. Qed. Lemma equivlistA_cons_nil x l : ¬equivlistA eqA (x :: l) []. Proof. intros E. now eapply InA_nil, E, InA_cons_hd. Qed. Lemma equivlistA_nil_eq l : equivlistA eqA l [] → l ≡ []. Proof. induction l. easy. intros E. edestruct equivlistA_cons_nil; eauto. Qed. Lemma InA_double_head z x l : InA eqA z (x :: x :: l) ↔ InA eqA z (x :: l). Proof. split; intros E1; auto. inversion_clear E1; auto. Qed. Lemma InA_permute_heads z x y l : InA eqA z (x :: y :: l) → InA eqA z (y :: x :: l). Proof. intros E1. inversion_clear E1 as [|?? E2]; auto. inversion_clear E2; auto. Qed. Lemma equivlistA_double_head x l : equivlistA eqA (x :: x :: l) (x :: l). Proof. split; apply InA_double_head. Qed. Lemma equivlistA_permute_heads x y l : equivlistA eqA (x :: y :: l) (y :: x :: l). Proof. split; apply InA_permute_heads. Qed. Lemma InA_app_ass z l₁ l₂ l₃ : InA eqA z (l₁ ++ (l₂ ++ l₃)) ↔ InA eqA z ((l₁ ++ l₂) ++ l₃). Proof. now rewrite app_ass. Qed. Lemma InA_app_nil_l z l : InA eqA z ([] ++ l) ↔ InA eqA z l. Proof. firstorder. Qed. Lemma InA_app_nil_r z l : InA eqA z (l ++ []) ↔ InA eqA z l. Proof. now rewrite app_nil_r. Qed. Lemma equivlistA_app_ass l₁ l₂ l₃ : equivlistA eqA (l₁ ++ (l₂ ++ l₃)) ((l₁ ++ l₂) ++ l₃). Proof. now rewrite app_ass. Qed. Lemma equivlistA_app_nil_l l : equivlistA eqA ([] ++ l) l. Proof. reflexivity. Qed. Lemma equivlistA_app_nil_r l : equivlistA eqA (l ++ []) l. Proof. now rewrite app_nil_r. Qed. Lemma InA_comm z l₁ l₂ : InA eqA z (l₁ ++ l₂) → InA eqA z (l₂ ++ l₁). Proof. rewrite !(InA_app_iff _). tauto. Qed. Lemma equivlistA_app_comm l₁ l₂ : equivlistA eqA (l₁ ++ l₂) (l₂ ++ l₁). Proof. split; apply InA_comm. Qed. Lemma InA_app_idem z l : InA eqA z (l ++ l) ↔ InA eqA z l. Proof. rewrite (InA_app_iff _). tauto. Qed. Lemma equivlistA_app_idem l : equivlistA eqA (l ++ l) l. Proof. split; apply InA_app_idem. Qed. Global Instance: Proper (equivlistA eqA ==> equivlistA eqA ==> equivlistA eqA) (@app A). Proof. intros ?? E1 ?? E2 ?. rewrite !(InA_app_iff _), E1, E2. tauto. Qed. Inductive PermutationA : list A → list A → Prop := | permA_nil: PermutationA [] [] | permA_skip x₁ x₂ l₁ l₂ : eqA x₁ x₂ → PermutationA l₁ l₂ → PermutationA (x₁ :: l₁) (x₂ :: l₂) | permA_swap x y l : PermutationA (y :: x :: l) (x :: y :: l) | permA_trans l₁ l₂ l₃ : PermutationA l₁ l₂ → PermutationA l₂ l₃ -> PermutationA l₁ l₃. Hint Constructors PermutationA. Global Instance: Equivalence PermutationA. Proof. split. intros l. induction l; intuition. intros l₁ l₂. induction 1; eauto. apply permA_skip; intuition. intros ???. now apply permA_trans. Qed. Global Instance PermutationA_cons : Proper (eqA ==> PermutationA ==> PermutationA) (@cons A). Proof. repeat intro. now apply permA_skip. Qed. Lemma PermutationA_app_head l₁ l₂ k : PermutationA l₁ l₂ → PermutationA (k ++ l₁) (k ++ l₂). Proof. intros. induction k. easy. apply permA_skip; intuition. Qed. Global Instance PermutationA_app : Proper (PermutationA ==> PermutationA ==> PermutationA) (@app A). Proof. intros l₁ l₂ Pl k₁ k₂ Pk. induction Pl. easy. now apply permA_skip. etransitivity. rewrite <-!app_comm_cons. now apply permA_swap. rewrite !app_comm_cons. now apply PermutationA_app_head. do 2 (etransitivity; try eassumption). apply PermutationA_app_head. now symmetry. Qed. Lemma PermutationA_app_tail l₁ l₂ k : PermutationA l₁ l₂ → PermutationA (l₁ ++ k) (l₂ ++ k). Proof. intros E. now rewrite E. Qed. Lemma PermutationA_cons_append l x : PermutationA (x :: l) (l ++ x :: nil). Proof. induction l. easy. simpl. rewrite <-IHl. intuition. Qed. Lemma PermutationA_app_comm l₁ l₂ : PermutationA (l₁ ++ l₂) (l₂ ++ l₁). Proof. induction l₁. now rewrite app_nil_r. rewrite <-app_comm_cons, IHl₁. now rewrite app_comm_cons, PermutationA_cons_append, <-app_assoc. Qed. Lemma PermutationA_cons_app l l₁ l₂ x : PermutationA l (l₁ ++ l₂) → PermutationA (x :: l) (l₁ ++ x :: l₂). Proof. intros E. rewrite E. now rewrite app_comm_cons, PermutationA_cons_append, <-app_assoc. Qed. Lemma PermutationA_middle l₁ l₂ x : PermutationA (x :: l₁ ++ l₂) (l₁ ++ x :: l₂). Proof. now apply PermutationA_cons_app. Qed. Lemma PermutationA_equivlistA l₁ l₂ : PermutationA l₁ l₂ → equivlistA eqA l₁ l₂. Proof. induction 1. reflexivity. now apply equivlistA_cons_proper. now apply equivlistA_permute_heads. etransitivity; eassumption. Qed. Lemma NoDupA_equivlistA_PermutationA l₁ l₂ : NoDupA eqA l₁ → NoDupA eqA l₂ → equivlistA eqA l₁ l₂ → PermutationA l₁ l₂. Proof. intros Pl₁. revert l₂. induction Pl₁ as [|x l₁ E1]. intros. now rewrite equivlistA_nil_eq by now symmetry. intros l₂ Pl₂ E2. destruct (@InA_split _ eqA l₂ x) as [l₂h [y [l₂t [E3 ?]]]]. rewrite <-E2. intuition. subst. transitivity (y :: l₁); [intuition |]. apply PermutationA_cons_app, IHPl₁. now apply NoDupA_split with y. apply equivlistA_NoDupA_split with x y; intuition. Qed. End equivlistA_misc. Arguments PermutationA {A} _ _ _. Notation "( x ::)" := (cons x) (only parsing) : mc_scope. Notation "(:: l )" := (λ x, cons x l) (only parsing) : mc_scope. Arguments app {A} _ _. Section contents. Context `{Setoid A}. Global Instance list_equiv: Equiv (list A) := eqlistA (=). Global Instance list_nil: MonUnit (list A) := []. Global Instance list_app: SgOp (list A) := app. Global Instance cons_proper: Proper (=) (@cons A). Proof. repeat intro. now apply eqlistA_cons. Qed. Instance: Setoid (list A). Proof. constructor; apply _. Qed. Instance app_proper: Proper ((=) ==> (=) ==> (=)) (@app A). Proof. apply _. Qed. Global Instance: Monoid (list A). Proof. repeat (split; try apply _). intros x y z. unfold sg_op, list_app. now rewrite app_ass. now repeat intro. intros x. change (x ++ [] = x). now rewrite app_nil_end. Qed. End contents. Lemma list_equiv_eq {A} (x y : list A) : @list_equiv A (≡) x y ↔ x ≡ y. Proof. split. induction 1. reflexivity. now f_equal. intros. now subst. Qed. #[global] Instance list_join: MonadJoin list := λ _, fix list_join ll := match ll with | [] => [] | l :: ll => l & list_join ll end. #[global] Instance list_map: SFmap list := map. #[global] Instance list_ret: MonadReturn list := λ _ x, [x]. #[global] Instance list_join_proper `{Setoid A} : Proper (=) (@list_join A). Proof. intros l. induction l; intros k E; inversion_clear E; try reflexivity. simpl. apply app_proper; auto. Qed. #[global] Instance list_ret_proper `{Equiv A}: Proper (=) (list_ret A). Proof. compute. firstorder. Qed. #[global] Instance list_map_proper `{Setoid A} `{Setoid B} : Proper (((=) ==> (=)) ==> ((=) ==> (=))) (list_map A B). Proof. intros f g E1 l. induction l; intros k E2; inversion_clear E2. reflexivity. simpl. apply cons_proper; auto. Qed. Lemma list_join_app `{Setoid A} (l k : list (list A)) : join (l & k) = join l & join k. Proof. unfold join, list_join, sg_op, list_app. induction l; simpl; try reflexivity. now rewrite IHl, app_assoc. Qed. Lemma list_map_app `{Equiv A} `{Equiv B} `{!Setoid_Morphism (f : A → B)} (l k : list A) : sfmap f (l & k) = sfmap f l & sfmap f k. Proof. pose proof (setoidmor_b f). now setoid_rewrite map_app. Qed. Local Instance: SFunctor list. Proof. split; try apply _; unfold sfmap, list_map, compose. intros. intros ???. now rewrite map_id. intros A ? B ? C ? f ? g ? ?? E. pose proof (setoidmor_a f). pose proof (setoidmor_b f). pose proof (setoidmor_a g). now rewrite <-E, map_map. Qed. #[global] Instance: StrongMonad list. Proof. split; try apply _; unfold compose, id, sfmap, join, ret. intros A ? B ? f [???]. intros ?? E. now rewrite <-E. intros A ? B ? f ? l k E. pose proof (setoidmor_a f). pose proof (setoidmor_b f). rewrite <-E. clear E. induction l; try reflexivity. simpl. setoid_rewrite <-IHl. now apply list_map_app. intros A ?? [|?] k E; inversion_clear E; try reflexivity. simpl. rewrite right_identity. now apply cons_proper. intros A ? ? l k E. rewrite <-E. clear E. induction l; try reflexivity. simpl. now rewrite IHl. intros A ?? l k E. rewrite <-E. clear E. induction l; try reflexivity. simpl. now rewrite IHl, list_join_app. Qed. #[global] Instance: FullMonad list. Proof. apply strong_monad_default_full_monad. Qed. (* (* so far so good, but now: *) Global Instance inj: InjectToSeq list := λ _ x, x :: nil. (* universe inconsistency... *) Global Instance sm: SeqToMonoid A (list A) := λ _ _ _ f, fold_right ((&) ∘ f) mon_unit. Instance: Equiv A := eq. Instance inj_mor: Setoid_Morphism inj. Section sm_mor. Context `{Monoid M} (f: A -> M). Lemma sm_app (a a': list A): sm M _ _ f (a ++ a') = sm M _ _ f a & sm M _ _ f a'. Proof with reflexivity. induction a; simpl; intros. rewrite monoid_lunit... unfold compose. rewrite IHa, associativity... Qed. Global Instance: Setoid_Morphism f → Monoid_Morphism (sm M _ _ f). Proof. intros. repeat (constructor; try apply _). intros. apply sm_app. reflexivity. Qed. End sm_mor. Global Instance: @Sequence A (list A) eq _ _ _ _ _. Proof with try reflexivity. apply (Build_Sequence A (list A) eq _ _ _ _ _ _ _ inj_mor _). unfold proves_free. simpl. split; repeat intro. destruct i. repeat (simpl in *; unfold compose). rewrite H. apply monoid_runit. destruct b. pose proof (H tt). clear H. rename H0 into H. simpl in *. change (∀ x y1 : A, x = y1 → ua.variety_equiv varieties.monoid.theory y tt (proj1_sig y0 tt (inj x)) (setoidcat.f _ _ (f tt) y1)) in H. revert a. destruct y0. simpl in *. pose proof (varieties.monoid.from_object y). assert (universal_algebra.Implementation varieties.monoid.sig (λ _ : universal_algebra.sorts varieties.monoid.sig, list A)). exact (ua.variety_op _ (varieties.monoid.object (list A))). pose proof (@varieties.monoid.morphism_from_ua (list A) _ y _ (ua.variety_op _ (varieties.monoid.object (list A))) (ua.variety_op _ y) x h _ _ tt). induction a. change (x tt mon_unit = mon_unit). apply preserves_mon_unit. replace (a :: a0) with ((a :: nil) & a0)... rewrite preserves_sg_op, IHa. rewrite (H a a)... Qed. (* todo: clean up *) *) math-classes-8.19.0/implementations/list_finite_set.v000066400000000000000000000245261460576051100227760ustar00rootroot00000000000000Require Import Coq.Lists.List Coq.Lists.SetoidList MathClasses.implementations.list MathClasses.interfaces.abstract_algebra MathClasses.interfaces.finite_sets MathClasses.interfaces.orders MathClasses.theory.lattices MathClasses.orders.lattices. Import ListNotations. (* We define finite sets as unordered lists. This implementation is slow, but quite convenient as a reference implementation to lift properties to arbitrary finite set instances. *) #[global] Instance listset A `{Equiv A} : SetType A | 30 := sig (NoDupA (=)). Section listset. Context `{Setoid A} `{∀ a₁ a₂ : A, Decision (a₁ = a₂)}. Instance listset_in_raw: Contains A (list A) := InA (=). Instance listset_equiv_raw: Equiv (list A) := equivlistA (=). Instance: Setoid (list A) := {}. Instance listset_empty_raw: Bottom (list A) := []. Instance listset_join_raw: Join (list A) := @app A. Instance: BoundedJoinSemiLattice (list A). Proof. split. split. split. split. apply _. repeat intro. now apply equivlistA_app_ass. apply _. repeat intro. now apply equivlistA_app_nil_l. repeat intro. now apply equivlistA_app_nil_r. repeat intro. now apply equivlistA_app_comm. repeat intro. now apply equivlistA_app_idem. Qed. Global Instance listset_to_list: Cast (set_type A) (list A) := @proj1_sig _ _. Global Instance listset_in: SetContains A := λ x l, x ∈ 'l. Global Instance listset_le: SetLe A := λ l k, ∀ x, x ∈ l → x ∈ k. Global Instance listset_equiv: SetEquiv A := λ l k, ∀ x, x ∈ l ↔ x ∈ k. Instance: Setoid (set_type A). Proof. now apply (setoids.projected_setoid listset_to_list). Qed. Global Instance: Setoid_Morphism listset_to_list. Proof. firstorder. Qed. Global Instance: Injective listset_to_list. Proof. firstorder. Qed. Global Instance: Proper ((=) ==> (=) ==> iff) listset_in. Proof. intros x y E1 l k E2. transitivity (listset_in x k). easy. unfold listset_in. now rewrite E1. Qed. Fixpoint listset_add_raw (x : A) (l : list A) : list A := match l with | [] => [x] | y :: l => y :: if decide_rel (=) x y then l else listset_add_raw x l end. Lemma listset_add_raw_cons l x : x :: l = listset_add_raw x l. Proof. induction l; simpl; try reflexivity. case (decide_rel _); intros E. now rewrite E, equivlistA_double_head. now rewrite equivlistA_permute_heads, IHl. Qed. Lemma listset_add_raw_InA (l : list A) (x y : A) : y ∈ listset_add_raw x l → y = x ∨ y ∈ l. Proof. unfold contains, listset_in_raw. induction l; simpl. intros E. inversion_clear E; auto. case (decide_rel _); auto; intros E1 E2. inversion_clear E2; intuition. Qed. Lemma listset_add_raw_NoDupA (l : list A) (x : A) : NoDupA (=) l → NoDupA (=) (listset_add_raw x l). Proof. intros Pl. induction l; simpl. now apply NoDupA_singleton. case (decide_rel _); intros E1; auto. inversion_clear Pl. apply NoDupA_cons; auto. intros E2. destruct (listset_add_raw_InA _ _ _ E2); intuition. Qed. Global Program Instance listset_empty: EmptySet A := []. Global Program Instance listset_singleton: SetSingleton A := λ x, [x]. Next Obligation. now apply NoDupA_singleton. Qed. Global Program Instance listset_join: SetJoin A := λ l k, fold_right listset_add_raw (`k) (`l)↾_. Next Obligation. destruct l as [l Pl], k as [k Pk]. induction l; intros; simpl in *; auto. apply listset_add_raw_NoDupA, IHl. now inversion Pl. Qed. Instance: Setoid_Morphism listset_singleton. Proof. split; try apply _. intros ? ? E. apply (injective listset_to_list). change ([x] = [y]). now rewrite E. Qed. Lemma listset_to_list_preserves_join l k : listset_to_list (l ⊔ k) = listset_to_list l ⊔ listset_to_list k. Proof. destruct l as [l Pl], k as [k Pk]. unfold join, listset_join, listset_join_raw. simpl. clear Pk Pl. induction l; simpl; intros; [easy|]. now rewrite <-IHl, listset_add_raw_cons. Qed. Instance: BoundedJoinSemiLattice (set_type A). Proof. apply (projected_bounded_sl listset_to_list). intros. now apply listset_to_list_preserves_join. reflexivity. Qed. Lemma listset_in_join l k x : x ∈ l ⊔ k ↔ x ∈ l ∨ x ∈ k. Proof. unfold contains, listset_in_raw, listset_in. rewrite listset_to_list_preserves_join. now apply InA_app_iff. Qed. Instance: JoinSemiLatticeOrder listset_le. Proof. apply alt_Build_JoinSemiLatticeOrder. intros l k. unfold le, listset_le, equiv, listset_equiv. setoid_rewrite listset_in_join. firstorder auto. Qed. Lemma listset_induction (P : set_type A → Prop) `{proper : !Proper ((=) ==> iff) P} : P ∅ → (∀ x l, x ∉ l → P l → P ({{ x }} ⊔ l)) → ∀ l, P l. Proof. intros Pempty Padd. intros [l Pl]. induction l as [|x l]. apply proper with ∅; firstorder. inversion_clear Pl as [|??? Pl']. apply proper with ({{ x }} ⊔ l↾Pl'); auto. intros z. change (z ∈ x :: l ↔ z ∈ listset_add_raw x l). now rewrite listset_add_raw_cons. Qed. Fixpoint listset_extend_raw `{Bottom B} `{Join B} (f : A → B) (l : list A) : B := match l with | [] => ⊥ | x :: l => f x ⊔ listset_extend_raw f l end. Global Instance list_extend: FSetExtend A := λ _ _ _ f l, listset_extend_raw f (`l). Section listset_extend. Context `{BoundedJoinSemiLattice B} `{!Setoid_Morphism (f : A → B)}. Lemma listset_extend_raw_permute (l k : list A) : PermutationA (=) l k → listset_extend_raw f l = listset_extend_raw f k. Proof. induction 1; simpl. reflexivity. apply sg_op_proper. now apply sm_proper. easy. now rewrite !associativity, (commutativity (f _)). etransitivity; eassumption. Qed. Instance list_extend_proper: Proper (equiv ==> equiv) (fset_extend f). Proof. intros [??][??] ?. apply listset_extend_raw_permute. now apply NoDupA_equivlistA_PermutationA. Qed. Lemma list_extend_empty: fset_extend f ∅ = ⊥. Proof. reflexivity. Qed. Lemma list_extend_add x l : fset_extend f ({{x}} ⊔ l) = f x ⊔ fset_extend f l. Proof. destruct l as [l Pl]. unfold fset_extend, list_extend. simpl. clear Pl. induction l; simpl; [easy|]. case (decide_rel _); intros E. now rewrite E, associativity, (idempotency (&) _). now rewrite IHl, 2!associativity, (commutativity (f _)). Qed. Instance list_extend_mor: BoundedJoinSemiLattice_Morphism (fset_extend f). Proof. repeat (split; try apply _). intros l k. change (fset_extend f (l ⊔ k) = fset_extend f l ⊔ fset_extend f k). pattern l. apply listset_induction; clear l. solve_proper. now rewrite list_extend_empty, 2!left_identity. intros x l E1 E2. now rewrite <-associativity, 2!list_extend_add, E2, associativity. reflexivity. Qed. End listset_extend. Local Existing Instance list_extend_mor. Global Instance: FSet A. Proof. split; try apply _. intros B ???? f ? x y E. unfold compose, fset_extend, list_extend. simpl. now rewrite E, right_identity. intros B ??? f ? h ? E1 k l E2. pose proof (bounded_join_slmor_b (f:=h)). rewrite E2. clear k E2. pattern l. apply listset_induction; clear l. solve_proper. now rewrite preserves_bottom. intros x l E2 E3. rewrite list_extend_add, preserves_join, E3. apply sg_op_proper; [|easy]. symmetry. now apply E1. Qed. Instance: FSetContainsSpec A. Proof. split; try apply _. unfold le, listset_le. intros x X; split; intros E1. intros z E2. inversion_clear E2 as [?? E3|?? E3]. now rewrite E3. now inversion E3. apply E1. now rapply InA_cons_hd. Qed. Instance listset_in_raw_dec: ∀ x (l : list A), Decision (x ∈ l) := λ x l, InA_dec (decide_rel (=)) x l. Global Instance listset_in_dec: ∀ x (l : set_type A), Decision (x ∈ l) := λ x l, InA_dec (decide_rel (=)) x ('l). Instance listset_meet_raw: Meet (list A) := fix listset_meet_raw l k := match l with | [] => [] | x :: l => if decide_rel (∈) x k then x :: listset_meet_raw l k else listset_meet_raw l k end. Lemma listset_in_meet_raw l k x : x ∈ l ⊓ k ↔ x ∈ l ∧ x ∈ k. Proof. unfold meet, contains, listset_in_raw. split. intros E; split; revert E. induction l; simpl. intuition. case (decide_rel); intros ? E; intuition. inversion_clear E; intuition. induction l; simpl. intros E1; inversion E1. case (decide_rel); intros ? E1; intuition. inversion_clear E1 as [?? E2|]; auto. now rewrite E2. intros [E1 E2]. induction l; simpl; [easy|]. case (decide_rel); intros E3. inversion_clear E1; intuition. inversion_clear E1 as [?? E4|]; intuition. destruct E3. now rewrite <-E4. Qed. Lemma listset_meet_raw_NoDupA (l k : list A) : NoDupA (=) l → NoDupA (=) (l ⊓ k). Proof. unfold meet. intros Pl. induction l; simpl; auto. inversion_clear Pl as [|? ? E1]. case (decide_rel); intros; auto. apply NoDupA_cons; auto. intros E2. destruct E1. now apply (listset_in_meet_raw l k _). Qed. Global Program Instance listset_meet: SetMeet A := λ l k, listset_meet_raw l k. Next Obligation. apply listset_meet_raw_NoDupA. now destruct l. Qed. Instance listset_diff_raw: Difference (list A) := fix listset_diff_raw l k := match l with | [] => [] | x :: l => if decide_rel (∈) x k then listset_diff_raw l k else x :: listset_diff_raw l k end. Lemma listset_in_diff_raw l k x : x ∈ l ∖ k ↔ x ∈ l ∧ x ∉ k. Proof. unfold difference, contains, listset_in_raw. split. intros E; split; revert E. induction l; simpl. intuition. case (decide_rel); intros ? E; intuition. inversion_clear E; intuition. induction l; simpl. intros E1; inversion E1. case (decide_rel); intros ? E1. intuition. inversion_clear E1 as [?? E2|]; auto. now rewrite E2. intros [E1 E2]. induction l; simpl; [easy|]. case (decide_rel); intros E3. inversion_clear E1 as [?? E4|]; intuition. destruct E2. now rewrite E4. inversion_clear E1; intuition. Qed. Lemma listset_diff_raw_NoDupA (l k : list A) : NoDupA (=) l → NoDupA (=) (l ∖ k). Proof. unfold difference. intros Pl. induction l; simpl; auto. inversion_clear Pl as [|? ? E1]. case (decide_rel); intros; auto. apply NoDupA_cons; auto. intros E2. destruct E1. now apply (listset_in_diff_raw l k _). Qed. Global Program Instance listset_diff: SetDifference A := λ l k, listset_diff_raw l k. Next Obligation. apply listset_diff_raw_NoDupA. now destruct l. Qed. Global Instance: FullFSet A | 30. Proof. split; try apply _. intros [??] [??]. now rapply listset_in_meet_raw. intros [??] [??]. now rapply listset_in_diff_raw. Qed. End listset. math-classes-8.19.0/implementations/modular_ring.v000066400000000000000000000016051460576051100222650ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.integers MathClasses.theory.integers MathClasses.theory.ring_ideals. Definition is_multiple `{Equiv Z} `{Mult Z} (b x : Z) := ∃ k, x = b * k. Notation Mod b := (Factor _ (is_multiple b)). Section modular_ring. Context `{Ring Z} {b : Z}. Add Ring R : (rings.stdlib_ring_theory Z). Global Instance: RingIdeal Z (is_multiple b). Proof. unfold is_multiple. split. solve_proper. split. exists 0, 0. ring. intros x y [k1 E1] [k2 E2]. exists (k1 - k2). rewrite E1, E2. ring. intros x y [k E]. exists (y * k). rewrite E. ring. intros x y [k E]. exists (x * k). rewrite E. ring. Qed. Lemma modular_ring_eq (x y : Mod b) : x = y ↔ ∃ k, 'x = 'y + b * k. Proof. split; intros [k E]; exists k. rewrite <-E. ring. rewrite E. ring. Qed. End modular_ring. math-classes-8.19.0/implementations/mset_finite_set.v000066400000000000000000000066661460576051100230000ustar00rootroot00000000000000Require Import Coq.MSets.MSetInterface Coq.MSets.MSetFacts Coq.MSets.MSetProperties MathClasses.implementations.list MathClasses.implementations.list_finite_set MathClasses.theory.finite_sets MathClasses.interfaces.finite_sets MathClasses.interfaces.orders MathClasses.interfaces.abstract_algebra. Import ListNotations. Module MSet_FSet (E : DecidableType) (Import set : WSetsOn E). Module facts := WFactsOn E set. Module props := WPropertiesOn E set. #[global] Instance mset: SetType elt := t. #[global] Instance mset_in: SetContains elt := In. #[global] Instance mset_car_eq: Equiv elt := E.eq. #[global] Instance mset_eq: SetEquiv elt := Equal. #[global] Instance mset_le: SetLe elt := Subset. #[global] Instance mset_singleton: SetSingleton elt := singleton. #[global] Instance mset_empty: EmptySet elt := empty. #[global] Instance mset_join: SetJoin elt := union. #[global] Instance mset_meet: SetMeet elt := inter. #[global] Instance mset_difference: SetDifference elt := diff. #[global] Instance mset_car_dec: ∀ x y : elt, Decision (x = y) := E.eq_dec. #[global] Instance mset_dec: ∀ x y : t, Decision (x = y) := eq_dec. Local Instance: Setoid elt. Proof. split; try apply _. Qed. Local Instance: BoundedJoinSemiLattice mset. Proof. Local Opaque Equal. repeat (split; try apply _); repeat intro. symmetry. now apply props.union_assoc. now apply props.empty_union_1, empty_spec. now apply props.empty_union_2, empty_spec. now apply props.union_sym. apply props.union_subset_equal, props.subset_refl. Local Transparent Equal. Qed. Local Instance: Setoid_Morphism singleton. Proof. split; try apply _. Qed. Definition to_listset (X : @set_type _ mset) : @set_type _ (listset elt) := props.to_list X↾elements_spec2w X. #[global] Instance from_listset: Inverse to_listset := λ l, props.of_list (`l). Local Instance: Setoid_Morphism to_listset. Proof. split; try apply _. intros X Y E x. change (InA E.eq x (props.to_list X) ↔ InA E.eq x (props.to_list Y)). now rewrite <-!props.of_list_1, 2!props.of_list_3, E. Qed. #[global] Instance: Bijective to_listset. Proof. split; split; try apply _. intros X Y E x. rewrite <-!elements_spec1. now apply E. intros ?? E. rewrite <-E. now rapply props.of_list_2. Qed. #[global] Instance: BoundedJoinSemiLattice_Morphism to_listset. Proof. split; try apply _; split; try apply _. split; try apply _. intros X Y z. setoid_rewrite listset_in_join. repeat setoid_rewrite elements_spec1. apply union_spec. intros z. compute. rewrite props.elements_empty. tauto. Qed. #[global] Instance mset_extend: FSetExtend elt := iso_is_fset_extend id to_listset. Local Instance: FSet elt. Proof. apply (iso_is_fset id to_listset). intros x y E z. change (InA E.eq z (elements (singleton x)) ↔ InA E.eq z [y]). now rewrite elements_spec1, singleton_spec, E, InA_singleton. Qed. #[global] Instance: FullFSet elt. Proof. split; try apply _. split. apply lattices.alt_Build_JoinSemiLatticeOrder. intros X Y. split; intros E. now apply props.union_subset_equal. rewrite <-E. now apply props.union_subset_1. intros x X. split; intros E. transitivity (add x empty). now apply props.subset_equal, props.singleton_equal_add. apply props.subset_add_3. auto. now apply props.subset_empty. apply props.union_subset_equal in E. rewrite <-E. now apply facts.union_2, singleton_spec. now apply inter_spec. now apply diff_spec. Qed. End MSet_FSet. math-classes-8.19.0/implementations/natpair_integers.v000066400000000000000000000115501460576051100231410ustar00rootroot00000000000000(* The standard library's implementation of the integers (BinInt) uses nasty binary positive crap with various horrible horrible bit fiddling operations on it (especially Pminus). The following is a much simpler implementation whose correctness can be shown much more easily. In particular, it lets us use initiality of the natural numbers to prove initiality of these integers. *) Require MathClasses.theory.naturals. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.theory.categories MathClasses.interfaces.naturals MathClasses.interfaces.integers MathClasses.theory.jections. Require Export MathClasses.implementations.semiring_pairs. Section contents. Context `{Naturals N}. Add Ring N : (rings.stdlib_semiring_theory N). Notation Z := (SRpair N). (* We show that Z is initial, and therefore a model of the integers. *) Global Instance SRpair_to_ring: IntegersToRing Z := λ _ _ _ _ _ _ z, naturals_to_semiring N _ (pos z) + - naturals_to_semiring N _ (neg z). (* Hint Rewrite preserves_0 preserves_1 preserves_mult preserves_plus: preservation. doesn't work for some reason, so we use: *) Ltac preservation := repeat (rewrite rings.preserves_plus || rewrite rings.preserves_mult || rewrite rings.preserves_0 || rewrite rings.preserves_1). Section for_another_ring. Context `{Ring R}. Add Ring R : (rings.stdlib_ring_theory R). Notation n_to_sr := (naturals_to_semiring N R). Notation z_to_r := (integers_to_ring Z R). Instance: Proper ((=) ==> (=)) z_to_r. Proof. intros [xp xn] [yp yn]. change (xp + yn = yp + xn → n_to_sr xp - n_to_sr xn = n_to_sr yp - n_to_sr yn). intros E. apply rings.equal_by_zero_sum. transitivity (n_to_sr xp + n_to_sr yn - (n_to_sr xn + n_to_sr yp)); [ring|]. rewrite <-2!rings.preserves_plus, E, (commutativity xn yp). ring. Qed. Ltac derive_preservation := unfold integers_to_ring, SRpair_to_ring; simpl; preservation; ring. Let preserves_plus x y: z_to_r (x + y) = z_to_r x + z_to_r y. Proof. derive_preservation. Qed. Let preserves_mult x y: z_to_r (x * y) = z_to_r x * z_to_r y. Proof. derive_preservation. Qed. Let preserves_1: z_to_r 1 = 1. Proof. derive_preservation. Qed. Let preserves_0: z_to_r 0 = 0. Proof. derive_preservation. Qed. Global Instance: SemiRing_Morphism z_to_r. Proof. repeat (split; try apply _). exact preserves_plus. exact preserves_0. exact preserves_mult. exact preserves_1. Qed. Section for_another_morphism. Context (f : Z → R) `{!SemiRing_Morphism f}. Definition g : N → R := f ∘ cast N (SRpair N). Instance: SemiRing_Morphism g. Proof. unfold g. repeat (split; try apply _). Qed. Lemma same_morphism: z_to_r = f. Proof. intros [p n] z' E. rewrite <- E. clear E z'. rewrite SRpair_splits. preservation. rewrite 2!rings.preserves_negate. now rewrite 2!(naturals.to_semiring_twice _ _ _). Qed. End for_another_morphism. End for_another_ring. Instance: Initial (rings.object Z). Proof. apply integer_initial. intros. now apply same_morphism. Qed. Global Instance: Integers Z := {}. Context `{!NatDistance N}. Global Program Instance SRpair_abs: IntAbs Z N := λ x, match nat_distance_sig (pos x) (neg x) with | inl (n↾E) => inr n | inr (n↾E) => inl n end. Next Obligation. rewrite <-(naturals.to_semiring_unique (cast N (SRpair N))). do 2 red. simpl. now rewrite rings.plus_0_r, commutativity. Qed. Next Obligation. rewrite <-(naturals.to_semiring_unique (cast N (SRpair N))). do 2 red. simpl. symmetry. now rewrite rings.plus_0_r, commutativity. Qed. Notation n_to_z := (naturals_to_semiring N Z). (* Without this opaque, typeclasses find a proof of Injective zero, from [id_injective] *) Typeclasses Opaque zero. Let zero_product_aux a b : n_to_z a * n_to_z b = 0 → n_to_z a = 0 ∨ n_to_z b = 0. Proof. rewrite <-rings.preserves_mult. rewrite <-(naturals.to_semiring_unique (SRpair_inject)). intros E. setoid_inject. assert (HN : (a = 0) ∨ (b = 0)) by now apply zero_product. destruct HN as [HN|HN]; [left|right]; rewrite HN; apply preserves_mon_unit. Qed. Global Instance: ZeroProduct Z. Proof. intros x y E. destruct (SRpair_abs x) as [[a A]|[a A]], (SRpair_abs y) as [[b B]|[b B]]. rewrite <-A, <-B in E |- *. now apply zero_product_aux. destruct (zero_product_aux a b) as [C|C]. rewrite A, B. now apply rings.negate_zero_prod_r. left. now rewrite <-A. right. apply rings.flip_negate_0. now rewrite <-B. destruct (zero_product_aux a b) as [C|C]. rewrite A, B. now apply rings.negate_zero_prod_l. left. apply rings.flip_negate_0. now rewrite <-A. right. now rewrite <-B. destruct (zero_product_aux a b) as [C|C]. now rewrite A, B, rings.negate_mult_negate. left. apply rings.flip_negate_0. now rewrite <-A. right. apply rings.flip_negate_0. now rewrite <-B. Qed. End contents. math-classes-8.19.0/implementations/ne_list.v000066400000000000000000000104151460576051100212370ustar00rootroot00000000000000(** This module should be [Require]d but not [Import]ed (except for the notations submodule). *) Require Import Coq.Unicode.Utf8 Coq.Lists.List Coq.Setoids.Setoid Coq.Classes.Morphisms Coq.Sorting.Permutation. #[global] Instance: ∀ A, Proper (@Permutation A ==> eq) (@length A). Proof Permutation_length. Section contents. Context {T: Type}. Inductive L: Type := one: T → L | cons: T → L → L. Fixpoint app (a b: L) {struct a}: L := match a with | one x => cons x b | cons x y => cons x (app y b) end. Fixpoint foldr {R} (o: T → R) (f: T → R → R) (a: L): R := match a with | one x => o x | cons x y => f x (foldr o f y) end. Fixpoint foldr1 (f: T → T → T) (a: L): T := match a with | one x => x | cons x y => f x (foldr1 f y) end. Definition head (l: L): T := match l with one x => x | cons x _ => x end. Fixpoint to_list (l: L): list T := match l with | one x => x :: nil | cons x xs => x :: to_list xs end. Local Coercion to_list: L >-> list. Fixpoint from_list (x: T) (xs: list T): L := match xs with | nil => one x | List.cons h t => cons x (from_list h t) end. Definition tail (l: L): list T := match l with one _ => nil | cons _ x => to_list x end. Lemma decomp_eq (l: L): l = from_list (head l) (tail l). Proof with auto. induction l... destruct l... simpl in *. rewrite IHl... Qed. Definition last: L → T := foldr1 (fun x y => y). Fixpoint replicate_Sn (x: T) (n: nat): L := match n with | 0 => one x | S n' => cons x (replicate_Sn x n') end. Fixpoint take (n: nat) (l: L): L := match l, n with | cons x xs, S n' => take n' xs | _, _ => one (head l) end. Lemma two_level_rect (P: L → Type) (Pone: ∀ x, P (one x)) (Ptwo: ∀ x y, P (cons x (one y))) (Pmore: ∀ x y z, P z → (∀ y', P (cons y' z)) → P (cons x (cons y z))): ∀ l, P l. Proof with auto. cut (∀ l, P l * ∀ x, P (cons x l)). intros. apply X. destruct l... revert t. induction l... intros. split. apply IHl. intro. apply Pmore; intros; apply IHl. Qed. Lemma tl_length (l: L): S (length (tl l)) = length l. Proof. destruct l; reflexivity. Qed. Notation ListPermutation := (@Permutation.Permutation _). Definition Permutation (x y: L): Prop := ListPermutation x y. Global Instance: Equivalence Permutation. Proof with intuition. unfold Permutation. split; repeat intro... transitivity y... Qed. Global Instance: Proper (Permutation ==> ListPermutation) to_list. Proof. firstorder. Qed. Lemma Permutation_ne_tl_length (x y: L): Permutation x y → length (tl x) = length (tl y). Proof. intro H. apply eq_add_S. do 2 rewrite tl_length. rewrite H. reflexivity. Qed. End contents. Module Export coercions. Coercion to_list: L >-> list. End coercions. Arguments L : clear implicits. Fixpoint tails {A} (l: L A): L (L A) := match l with | one x => one (one x) | cons x y => cons l (tails y) end. Lemma tails_are_shorter {A} (y x: L A): In x (tails y) → length x <= length y. Proof with auto. induction y; simpl. intros [[] | ?]; intuition. intros [[] | C]... Qed. Fixpoint map {A B} (f: A → B) (l: L A): L B := match l with | one x => one (f x) | cons h t => cons (f h) (map f t) end. Lemma list_map {A B} (f: A → B) (l: L A): to_list (map f l) = List.map f (to_list l). Proof. induction l. reflexivity. simpl. congruence. Qed. Global Instance: forall {A B} (f: A → B), Proper (Permutation ==> Permutation) (map f). Proof with auto. intros ????? E. unfold Permutation. do 2 rewrite list_map. rewrite E. reflexivity. Qed. Fixpoint inits {A} (l: L A): L (L A) := match l with | one x => one (one x) | cons h t => cons (one h) (map (cons h) (inits t)) end. Module notations. Global Notation ne_list := L. Global Infix ":::" := cons (at level 60, right associativity). (* Todo: Try to get that "[ x ; .. ; y ]" notation working. *) Fixpoint ne_zip {A B: Type} (l: ne_list A) (m: ne_list B) {struct l} : ne_list (A * B) := match l with | one a => one (a, head m) | a ::: l => match m with | one b => one (a, b) | b ::: m => (a, b) ::: ne_zip l m end end. End notations. math-classes-8.19.0/implementations/nonneg_integers_naturals.v000066400000000000000000000056541460576051100247100ustar00rootroot00000000000000Require MathClasses.implementations.peano_naturals. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.integers MathClasses.interfaces.naturals MathClasses.interfaces.orders MathClasses.interfaces.additional_operations MathClasses.theory.int_abs. Require Export MathClasses.implementations.nonneg_semiring_elements. Section nonneg_integers_naturals. Context `{Integers Z} `{Apart Z} `{!TrivialApart Z} `{!FullPseudoSemiRingOrder Zle Zlt}. Add Ring Z: (rings.stdlib_ring_theory Z). (* We show that [Z⁺] is an instance of the naturals by constructing a retract to [nat] *) Program Let of_nat (x : nat) : Z⁺ := (naturals_to_semiring nat Z x)↾_. Next Obligation. apply nat_int.to_semiring_nonneg. Qed. Local Ltac unfold_equivs := unfold equiv, sig_equiv in *; simpl in *. Instance: Proper ((=) ==> (=)) of_nat. Proof. intros ?? E. unfold_equivs. now rewrite E. Qed. Instance: Setoid_Morphism of_nat := {}. Instance: SemiRing_Morphism of_nat. Proof. repeat (split; try apply _); repeat intro; unfold_equivs. now apply rings.preserves_plus. unfold mon_unit, zero_is_mon_unit. now apply rings.preserves_0. now apply rings.preserves_mult. unfold mon_unit, one_is_mon_unit. now apply rings.preserves_1. Qed. Program Let to_nat: Inverse of_nat := λ x, int_abs Z nat (`x). Existing Instance to_nat. Instance: Proper ((=) ==> (=)) to_nat. Proof. intros [x Ex] [y Ey] E. unfold to_nat. do 2 red in E. simpl in E. simpl. now rewrite E. Qed. Instance: SemiRing_Morphism to_nat. Proof. pose proof (_ : SemiRing (Z⁺)). repeat (split; try apply _). intros [x Ex] [y Ey]. unfold to_nat; unfold_equivs. simpl. now apply int_abs_nonneg_plus. unfold mon_unit, zero_is_mon_unit. now apply int_abs_0. intros [x Ex] [y Ey]. unfold to_nat; unfold_equivs. simpl. now apply int_abs_mult. unfold mon_unit, one_is_mon_unit. apply int_abs_1. Qed. Instance: Surjective of_nat. Proof. split. 2: apply _. intros [x Ex] y E. rewrite <- E. unfold to_nat, of_nat. unfold_equivs. now apply int_abs_nonneg. Qed. Global Instance: NaturalsToSemiRing (Z⁺) := naturals.retract_is_nat_to_sr of_nat. Global Instance: Naturals (Z⁺) := naturals.retract_is_nat of_nat. Context `{∀ x y : Z, Decision (x ≤ y)}. Global Program Instance ZPos_cut_minus: CutMinus (Z⁺) := λ x y, if decide_rel (≤) x y then 0 else ((x : Z) - (y : Z))↾_. Next Obligation. apply <-rings.flip_nonneg_minus. now apply orders.le_flip. Qed. Global Instance: CutMinusSpec (Z⁺) ZPos_cut_minus. Proof. split; intros [x Ex] [y Ey] E1; unfold cut_minus, ZPos_cut_minus; unfold_equivs. case (decide_rel); intros E2. rewrite left_identity. now apply (antisymmetry (≤)). simpl. ring. case (decide_rel); intros E2. reflexivity. simpl. apply (antisymmetry (≤)). now apply rings.flip_nonpos_minus. now apply rings.flip_nonneg_minus, orders.le_flip. Qed. End nonneg_integers_naturals. math-classes-8.19.0/implementations/nonneg_semiring_elements.v000066400000000000000000000072331460576051100246630ustar00rootroot00000000000000Require MathClasses.theory.rings. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.orders.semirings. Local Existing Instance pseudo_srorder_semiring. Section nonneg_semiring_elements. Context `{FullPseudoSemiRingOrder R} `{Apart R}. Add Ring R : (rings.stdlib_semiring_theory R). (* * Embedding of R⁺ into R *) Global Instance NonNeg_inject: Cast (R⁺) R := @proj1_sig R _. (* Operations *) Global Program Instance NonNeg_plus: Plus (R⁺) := λ x y, (`x + `y)↾_. Next Obligation. destruct x as [x Hx], y as [y Hy]. now apply nonneg_plus_compat. Qed. Global Program Instance NonNeg_mult: Mult (R⁺) := λ x y, (`x * `y)↾_. Next Obligation. destruct x as [x Hx], y as [y Hy]. now apply nonneg_mult_compat. Qed. Global Program Instance NonNeg_0: Zero (R⁺) := 0↾_. Global Program Instance NonNeg_1: One (R⁺) := 1↾_. Next Obligation. apply le_0_1. Qed. (* * Equalitity *) Local Ltac unfold_equivs := unfold equiv, sig_equiv in *; simpl in *. Instance: Proper ((=) ==> (=) ==> (=)) NonNeg_plus. Proof. intros [x1 Ex1] [y1 Ey1] E1 [x2 Ex2] [y2 Ey2] E2. unfold_equivs. now rewrite E1, E2. Qed. Instance: Proper ((=) ==> (=) ==> (=)) NonNeg_mult. Proof. intros [x1 Ex1] [y1 Ey1] E1 [x2 Ex2] [y2 Ey2] E2. unfold_equivs. now rewrite E1, E2. Qed. (* It is indeed a semiring *) Global Instance: SemiRing (R⁺). Proof. repeat (split; try apply _); repeat intro; unfold_equivs; ring. Qed. Instance: Proper ((=) ==> (=)) NonNeg_inject. Proof. now repeat intro. Qed. Global Instance: SemiRing_Morphism NonNeg_inject. Proof. repeat (split; try apply _); now repeat intro. Qed. Global Instance: Injective NonNeg_inject. Proof. split. trivial. apply _. Qed. (* Misc *) Global Instance NonNeg_trivial_apart `{!TrivialApart R} : TrivialApart (R⁺). Proof. intros x y. now Tactics.rapply trivial_apart. Qed. Global Instance NonNeg_equiv_dec `{∀ x y : R, Decision (x = y)} : ∀ x y: R⁺, Decision (x = y) := λ x y, decide_rel (=) ('x : R) ('y : R). Global Instance NonNeg_apart_dec `{∀ x y : R, Decision (x ≶ y)} : ∀ x y: R⁺, Decision (x ≶ y) := λ x y, decide_rel (≶) ('x : R) ('y : R). (* Order *) Global Instance NonNeg_le: Le (R⁺) := λ x y, 'x ≤ 'y. Global Instance NonNeg_lt: Lt (R⁺) := λ x y, 'x < 'y. Global Instance: Proper ((=) ==> (=) ==> iff) NonNeg_le. Proof. intros x1 y1 E1 x2 y2 E2. unfold NonNeg_le. now rewrite E1, E2. Qed. Instance: PartialOrder NonNeg_le. Proof. now apply (maps.projected_partial_order NonNeg_inject). Qed. Global Instance: OrderEmbedding NonNeg_inject. Proof. repeat (split; try apply _); intuition. Qed. (* Global Instance: TotalOrder NonNeg_le. Proof maps.embed_totalorder NonNeg_le. Global Instance: SemiRingOrder NonNeg_order. Proof. split; try apply _. intros x y. split; intros E. apply (order_preserving NonNeg_inject) in E. apply srorder_plus in E. destruct E as [z [Ez1 Ez2]]. exists (z↾Ez1); intuition. destruct E as [z [Ez1 Ez2]]. rewrite Ez2. apply (order_reflecting NonNeg_inject). rewrite rings.preserves_plus. now apply nonneg_plus_compat_r. intros x E1 y E2. apply (order_reflecting NonNeg_inject). rewrite rings.preserves_0, rings.preserves_mult. now apply srorder_mult. Qed. *) Global Program Instance NonNeg_le_dec `{∀ x y : R, Decision (x ≤ y)} : ∀ x y: R⁺, Decision (x ≤ y) := λ x y, match decide_rel (≤) ('x) ('y) with | left E => left _ | right E => right _ end. End nonneg_semiring_elements. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Opaque NonNeg_le. #[global] Typeclasses Opaque NonNeg_lt. math-classes-8.19.0/implementations/nonzero_field_elements.v000066400000000000000000000030371460576051100243350ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.theory.fields. (* The non zero elements of a field form a CommutativeMonoid. *) Section contents. Context `{Field F}. Add Ring F : (rings.stdlib_ring_theory F). Global Program Instance NonZero_1: One (F ₀) := (1:F). Next Obligation. solve_propholds. Qed. Global Program Instance NonZero_mult: Mult (F ₀) := λ x y, (`x * `y : F). Next Obligation. solve_propholds. Qed. Ltac unfold_equiv := repeat intro; unfold equiv, sig_equiv in *; simpl in *. Instance: Proper ((=) ==> (=) ==> (=)) NonZero_mult. Proof. intros [??] [??] E1 [??] [??] E2. unfold_equiv. now rewrite E1, E2. Qed. Global Instance: CommutativeMonoid (F ₀). Proof. repeat (split; try apply _); red; unfold_equiv. now apply associativity. now apply left_identity. now apply right_identity. now apply commutativity. Qed. Definition NonZero_inject (x : F ₀) : F := `x. Global Instance: SemiGroup_Morphism NonZero_inject (Bop:=(.*.)). Proof. repeat (split; try apply _); now repeat intro. Qed. Lemma quotients (a c : F) (b d : F ₀) : a // b + c // d = (a * `d + c * `b) // (b * d). Proof. setoid_replace (a // b) with ((a * `d) // (b * d)) by (apply equal_quotients; simpl; ring). setoid_replace (c // d) with ((`b * c) // (b * d)) by (apply equal_quotients; simpl; ring). ring. Qed. Lemma recip_distr `{∀ z, PropHolds (z ≠ 0) → LeftCancellation (.*.) z} (x y : F ₀) : // (x * y) = // x * // y. Proof. destruct x, y. apply recip_distr_alt. Qed. End contents. math-classes-8.19.0/implementations/option.v000066400000000000000000000100561460576051100211130ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.monads MathClasses.theory.jections MathClasses.theory.monads. Inductive option_equiv A `{Equiv A} : Equiv (option A) := | option_equiv_Some : Proper ((=) ==> (=)) Some | option_equiv_None : None = None. #[global] Existing Instance option_equiv. #[global] Hint Constructors option_equiv. Section contents. Context `{Setoid A}. Lemma Some_ne_None x : Some x ≠ None. Proof. intros E. inversion E. Qed. Lemma None_ne_Some x : None ≠ Some x. Proof. intros E. inversion E. Qed. Global Instance: Setoid (option A). Proof. split. intros [x|]. now apply option_equiv_Some. now apply option_equiv_None. red. induction 1. now apply option_equiv_Some. now apply option_equiv_None. intros x y z E. revert z. induction E. intros z E2. inversion_clear E2. apply option_equiv_Some. etransitivity; eassumption. easy. Qed. Global Instance: Setoid_Morphism Some. Proof. split; try apply _. now apply option_equiv_Some. Qed. Global Instance: Injective Some. Proof. split; try apply _. intros ? ? E. now inversion E. Qed. Global Instance option_zero: MonUnit (option A) := None. Global Instance option_op: SgOp (option A) := λ mx my, match mx with | Some x => Some x | None => my end. Instance: Proper ((=) ==> (=) ==> (=)) option_op. Proof. intros [?|] [?|] ? ???; simpl; try setoid_discriminate; auto. Qed. Global Instance: Monoid (option A). Proof. repeat (split; try apply _); now repeat intros [?|]. Qed. Lemma option_equiv_alt x y : x = y ↔ (∀ n, x = Some n ↔ y = Some n). Proof. split; intros E1. intro. now rewrite E1. destruct x, y. now apply E1. symmetry. now apply E1. now apply E1. easy. Qed. Global Program Instance option_dec `(A_dec : ∀ x y : A, Decision (x = y)) : ∀ x y : option A, Decision (x = y) := λ x y, match x with | Some r => match y with | Some s => match A_dec r s with left _ => left _ | right _ => right _ end | None => right _ end | None => match y with | Some s => right _ | None => left _ end end. Next Obligation. now apply (injective_ne Some). Qed. Next Obligation. setoid_discriminate. Qed. Next Obligation. setoid_discriminate. Qed. End contents. #[global] Hint Extern 10 (Equiv (option _)) => apply @option_equiv : typeclass_instances. Lemma option_equiv_eq {A} (x y : option A) : @option_equiv A (≡) x y ↔ x ≡ y. Proof. split. induction 1. now f_equal. reflexivity. intros. subst. now assert (@Setoid A (@eq A)) by (split; apply _). Qed. #[global] Instance option_ret: MonadReturn option := λ A x, Some x. #[global] Instance option_bind: MonadBind option := λ A B f x, match x with | Some a => f a | None => None end. #[global] Instance option_ret_proper `{Equiv A} : Proper ((=) ==> (=)) (option_ret A). Proof. intros x y E. now apply option_equiv_Some. Qed. #[global] Instance option_bind_proper `{Setoid A} `{Setoid (option B)}: Proper (=) (option_bind A B). Proof. intros f₁ f₂ E1 x₁ x₂ [?|]. unfold option_bind. simpl. now apply E1. easy. Qed. #[global] Instance: Monad option. Proof. repeat (split; try apply _); unfold bind, ret, option_bind, option_ret, compose. intros. now apply setoids.ext_equiv_refl. now intros ? ? ? [?|]. intros A ? B ? C ? ? f [???] g [???] [x|] [y|] E; try solve [inversion_clear E]. setoid_inject. cut (g x = g y); [|now rewrite E]. case (g x), (g y); intros E2; inversion_clear E2. now f_equiv. easy. easy. Qed. #[global] Instance option_map: SFmap option := option_map. #[global] Instance: FullMonad option. Proof. apply @monad_default_full_monad. apply _. Qed. Section map. Context `{Setoid A} `{Setoid B} `{!Injective (f : A → B)}. Global Instance: Injective (sfmap f). Proof. pose proof (injective_mor f). repeat (split; try apply _). intros [x|] [y|] E; try solve [inversion E | easy]. apply sm_proper. apply (injective f). now apply (injective Some). Qed. End map. math-classes-8.19.0/implementations/peano_naturals.v000066400000000000000000000112651460576051100226210ustar00rootroot00000000000000Require MathClasses.theory.ua_homomorphisms. Require Import Coq.Classes.Morphisms Coq.setoid_ring.Ring Coq.Arith.Arith_base Coq.Arith.PeanoNat MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals MathClasses.theory.categories MathClasses.interfaces.additional_operations MathClasses.interfaces.orders MathClasses.orders.semirings. #[global] Instance nat_equiv: Equiv nat := eq. #[global] Instance nat_plus: Plus nat := Peano.plus. #[global] Instance nat_0: Zero nat := 0%nat. #[global] Instance nat_1: One nat := 1%nat. #[global] Instance nat_mult: Mult nat := Peano.mult. #[global] Instance: SemiRing nat. Proof. repeat (split; try apply _); repeat intro. now apply Nat.add_assoc. now apply Nat.add_0_r. now apply Nat.add_comm. now apply Nat.mul_assoc. now apply Nat.mul_1_l. now apply Nat.mul_1_r. now apply Nat.mul_comm. now apply Nat.mul_add_distr_l. Qed. (* misc *) #[global] Instance: Injective S. Proof. repeat (split; try apply _). intros ?? E. now injection E. Qed. Global Instance nat_dec: ∀ x y: nat, Decision (x = y) := eq_nat_dec. Add Ring nat: (rings.stdlib_semiring_theory nat). Close Scope nat_scope. #[global] Instance: NaturalsToSemiRing nat := λ _ _ _ _ _, fix f (n: nat) := match n with 0%nat => 0 | S n' => f n' + 1 end. Section for_another_semiring. Context `{SemiRing R}. Notation toR := (naturals_to_semiring nat R). Add Ring R: (rings.stdlib_semiring_theory R). Instance: Proper ((=) ==> (=)) toR. Proof. unfold equiv, nat_equiv. repeat intro. subst. reflexivity. Qed. Let f_preserves_0: toR 0 = 0. Proof. reflexivity. Qed. Let f_preserves_1: toR 1 = 1. Proof. unfold naturals_to_semiring. simpl. ring. Qed. Let f_preserves_plus a a': toR (a + a') = toR a + toR a'. Proof with ring. induction a. change (toR a' = 0 + toR a')... change (toR (a + a') + 1 = toR (a) + 1 + toR a'). rewrite IHa... Qed. Let f_preserves_mult a a': toR (a * a') = toR a * toR a'. Proof with ring. induction a. change (0 = 0 * toR a')... change (toR (a' + a * a') = (toR a + 1) * toR a'). rewrite f_preserves_plus, IHa... Qed. Global Instance: SemiRing_Morphism (naturals_to_semiring nat R). repeat (constructor; try apply _). apply f_preserves_plus. apply f_preserves_0. apply f_preserves_mult. apply f_preserves_1. Qed. End for_another_semiring. Lemma O_nat_0 : O ≡ 0. Proof. reflexivity. Qed. Lemma S_nat_plus_1 x : S x ≡ x + 1. Proof. rewrite commutativity. reflexivity. Qed. Lemma S_nat_1_plus x : S x ≡ 1 + x. Proof. reflexivity. Qed. Lemma nat_induction (P : nat → Prop) : P 0 → (∀ n, P n → P (1 + n)) → ∀ n, P n. Proof nat_ind P. #[global] Instance: Initial (semirings.object nat). Proof. intros. apply natural_initial. intros. intros x y E. unfold equiv, nat_equiv in E. subst y. induction x. replace 0%nat with (zero:nat) by reflexivity. rewrite 2!rings.preserves_0. reflexivity. rewrite S_nat_1_plus. rewrite 2!rings.preserves_plus, 2!rings.preserves_1. now rewrite IHx. Qed. (* [nat] is indeed a model of the naturals *) #[global] Instance: Naturals nat := {}. (* Misc *) #[global] Instance: NoZeroDivisors nat. Proof. now intros x [Ex [y [Ey1 [H | H]%Nat.eq_mul_0]]]. Qed. #[global] Instance: ∀ z : nat, PropHolds (z ≠ 0) → LeftCancellation (.*.) z. Proof. intros z Ez x y. now apply Nat.mul_cancel_l. Qed. (* Order *) #[global] Instance nat_le: Le nat := Peano.le. #[global] Instance nat_lt: Lt nat := Peano.lt. #[global] Instance: FullPseudoSemiRingOrder nat_le nat_lt. Proof. assert (TotalRelation nat_le). intros x y. now destruct (le_ge_dec x y); intuition. assert (PartialOrder nat_le). split; try apply _. intros x y E. now apply Nat.le_antisymm. assert (SemiRingOrder nat_le). repeat (split; try apply _). intros x y E. exists (y - x)%nat. now rewrite Nat.add_comm, Nat.sub_add by (exact E). intros. now apply Nat.add_le_mono_l. intros. now apply Nat.add_le_mono_l with z. intros x y ? ?. change (0 * 0 <= x * y)%nat. now apply Nat.mul_le_mono. apply dec_full_pseudo_srorder. now apply Nat.le_neq. Qed. #[global] Instance: OrderEmbedding S. Proof. repeat (split; try apply _). exact le_n_S. exact le_S_n. Qed. #[global] Instance: StrictOrderEmbedding S. Proof. split; try apply _. Qed. #[global] Instance nat_le_dec : `{Decision (x ≤ y)} := le_dec. #[global] Instance nat_cut_minus: CutMinus nat := minus. #[global] Instance: CutMinusSpec nat nat_cut_minus. Proof. split. symmetry. rewrite commutativity. now rewrite Nat.add_comm, Nat.sub_add by (exact H). intros x y E. destruct (orders.le_equiv_lt x y E) as [E2|E2]. rewrite E2. now apply Nat.sub_diag. apply Nat.sub_0_le; exact E. Qed. math-classes-8.19.0/implementations/polynomials.v000066400000000000000000000265231460576051100221570ustar00rootroot00000000000000Require Import Coq.Lists.List MathClasses.interfaces.abstract_algebra MathClasses.interfaces.vectorspace MathClasses.theory.rings Ring. Import ListNotations. Section contents. Context R `{Ring R}. Add Ring R: (stdlib_ring_theory R). Definition poly := list R. Coercion poly_constant (c : R) : poly := [c]. Global Instance poly_zero: Zero poly := []. Global Instance poly_one: One poly := poly_constant 1. Definition all (l: list Prop): Prop := fold_left and l True. Lemma all_cons P Ps: all (P::Ps) ↔ P ∧ all Ps. Proof. unfold all. revert P; generalize True as Q. induction Ps as [|P' Ps IH]; intros. { cbn; tauto. } change (fold_left and (P' :: Ps) (Q ∧ P) ↔ P ∧ fold_left and (P'::Ps) Q). rewrite !IH. transitivity (P' ∧ (P ∧ fold_left and Ps Q)); try tauto. now rewrite <- (IH Q P). Qed. Arguments all: simpl never. Definition poly_eq_zero: poly → Prop := all ∘ map ((=) 0). Corollary poly_eq_zero_cons x p: poly_eq_zero (x::p) ↔ x = 0 ∧ poly_eq_zero p. Proof. unfold poly_eq_zero, compose; cbn; rewrite all_cons. intuition now symmetry. Qed. Lemma poly_eq_zero_ind (P: poly → Prop) (case_nil: P []) (casecons: ∀ x p, x = 0 → poly_eq_zero p → P p → P (x::p)) p: poly_eq_zero p → P p. Proof. induction p as [|x p IH]; auto. intros [??]%poly_eq_zero_cons. eauto. Qed. Global Instance poly_eq: Equiv poly := fix F p q := match p, q with | [], _ => poly_eq_zero q | _, [] => poly_eq_zero p | h :: t, h' :: t' => h = h' ∧ F t t' end. Lemma poly_eq_p_zero p: (p = 0) ↔ poly_eq_zero p. Proof. now destruct p. Qed. Instance: Reflexive poly_eq. Proof with intuition. repeat intro. induction x... split... Qed. Lemma poly_eq_cons : ∀ (a b : R) (p q : poly), (a = b /\ poly_eq p q) <-> poly_eq (a :: p) (b :: q). Proof. easy. Qed. Lemma poly_eq_ind (P: poly → poly → Prop) (case_0: ∀ p p', poly_eq_zero p → poly_eq_zero p' → P p p') (case_cons: ∀ x x' p p', x = x' → p = p' → P p p' → P (x::p) (x'::p')) p: ∀ p', p = p' → P p p'. Proof. induction p as [|x p IH]; intros [|x' p'] eqxy. 1,2,3: now apply case_0. destruct eqxy. apply case_cons; auto. Qed. Lemma poly_eq_zero_trans p q: poly_eq_zero p → poly_eq_zero q → p = q. Proof. revert q. induction p as [|x p IH]; intros ? eqp eqq; auto. destruct q as [|y q]; auto. rewrite poly_eq_zero_cons in *. rewrite <- poly_eq_cons. destruct eqp as [-> ?], eqq as [-> ?]; split; eauto. now apply IH. Qed. Instance: Symmetric poly_eq. Proof. intros p p' eqp. pattern p, p'; apply poly_eq_ind; auto; clear p p' eqp. - now intros; apply poly_eq_zero_trans. - intros ???? eqx eqp IH. rewrite <- poly_eq_cons; auto. Qed. Instance poly_eq_zero_proper: Proper (poly_eq ==> iff) poly_eq_zero. Proof. apply proper_sym_impl_iff. { apply _. } red; refine (poly_eq_ind _ _ _). - intros; hnf; auto. - unfold impl; intros ???? eqx eqp IH. rewrite !poly_eq_zero_cons, eqx; tauto. Qed. Instance: Transitive poly_eq. Proof. intros ??? eqxy; revert z. pattern x, y; refine (poly_eq_ind _ _ _ x y eqxy); clear x y eqxy. { intros x y x0 y0 z eqz. apply poly_eq_zero_trans; auto. now rewrite <- eqz. } intros ???? eqx eqp IH z eqz. destruct z as [|x'' p'']. { eapply poly_eq_zero_proper; eauto. now split. } destruct eqz as [eqx' eqp']; split; eauto. Qed. Global Instance: Setoid poly. Proof. split; try apply _. Qed. Global Instance: Plus poly := fix F p q := match p, q with | [], _ => q | _, [] => p | h :: t, h' :: t' => h + h' :: F t t' end. Lemma poly_eq_zero_plus_l p q: poly_eq_zero p → p + q = q. Proof. intro eqp; revert q. induction eqp as [|x p eqx eqp IH] using poly_eq_zero_ind. { easy. } intros [|y q]. { cbn -[poly_eq_zero]. rewrite poly_eq_zero_cons; auto. } cbn; split; auto. ring [eqx]. Qed. Instance plus_commutative: Commutative (+). Proof with (try easy); cbn. intro. induction x as [|x p IH]; intros [|y q]... split; auto; ring. Qed. Corollary poly_eq_zero_plus_r p q: poly_eq_zero q → p + q = p. Proof. rewrite commutativity. apply poly_eq_zero_plus_l. Qed. Corollary poly_eq_zero_plus p q: poly_eq_zero p → poly_eq_zero q → poly_eq_zero (p+q). Proof. intro. rewrite <- !poly_eq_p_zero; intro. now rewrite poly_eq_zero_plus_l. Qed. Global Instance poly_plus_proper: Proper (poly_eq ==> poly_eq ==> poly_eq) (+). Proof. unfold Proper, respectful. refine (poly_eq_ind _ _ _). { intros p p' zp zp' q q' eqq. rewrite !poly_eq_zero_plus_l; auto. } intros ???? eqx eqp IH. refine (poly_eq_ind _ _ _). { intros q q' zq zq'. rewrite !poly_eq_zero_plus_r; auto. cbn; auto. } intros y y' q q' eqy eqq _. cbn; split; eauto. ring [eqx eqy]. Qed. Instance plus_associative: Associative (+). Proof with try easy. do 2 red; induction x as [|x p IH]... intros [|y q]... intros [|z r]... cbn; split; auto. ring. Qed. Instance plus_left_id: LeftIdentity (+) 0. Proof. now intro; rewrite poly_eq_zero_plus_l. Qed. Instance plus_right_id: RightIdentity (+) 0. Proof. now intro; rewrite poly_eq_zero_plus_r. Qed. Instance poly_plus_monoid: Monoid poly. Proof. repeat (split; try apply _). Qed. Global Instance: Negate poly := map (-). Lemma poly_negate_zero p: poly_eq_zero p ↔ poly_eq_zero (-p). Proof. induction p as [|x p IH]. { easy. } cbn. rewrite !poly_eq_zero_cons, IH. enough (x = 0 ↔ -x = 0) by tauto. split; intro eq0; ring [eq0]. Qed. Instance poly_negate_proper: Proper (poly_eq ==> poly_eq) (-). Proof. refine (poly_eq_ind _ _ _). { now intros ?? ->%poly_negate_zero%poly_eq_p_zero ->%poly_negate_zero%poly_eq_p_zero. } intros ???? eqx eqp IH. cbn; split; eauto. Qed. Instance poly_negate_l: LeftInverse (+) (-) 0. Proof. intro; rewrite poly_eq_p_zero. induction x as [|x p IH]; cbn. { easy. } rewrite poly_eq_zero_cons; split; auto. ring. Qed. Instance poly_negate_r: RightInverse (+) (-) 0. Proof. now intro; rewrite commutativity, left_inverse. Qed. Instance poly_plus_abgroup: AbGroup poly. Proof. repeat (split; try apply _). Qed. Global Instance: ScalarMult R poly := fix F c q := match q with | [] => 0 | d :: q1 => c*d :: F c q1 end. Lemma poly_scalar_mult_0_r q c: poly_eq_zero q → poly_eq_zero (c · q). Proof. induction q as [|x q IH]. { easy. } cbn. rewrite !poly_eq_zero_cons. intros [-> ?]; split; auto. ring. Qed. Instance poly_scalar_mult_proper: Proper ((=) ==> (=) ==> (=)) (·). Proof. intros c c' eqc p p' eqp. revert p p' eqp; refine (poly_eq_ind _ _ _). { now intros; apply poly_eq_zero_trans; apply poly_scalar_mult_0_r. } intros ???? eqx eqp IH. split; auto; cbn. ring [eqx eqc]. Qed. Lemma poly_scalar_mult_1_r x: x · 1 = [x]. Proof. cbn; split; [ring|easy]. Qed. Instance poly_scalar_mult_1_l: LeftIdentity (·) 1. Proof. red; induction y as [|y p IH]; [easy|cbn]. split; auto; ring. Qed. Instance poly_scalar_mult_dist_l: LeftHeteroDistribute (·) (+) (+). Proof. intros a p. induction p as [|x p IH]; intros [|y q]; [easy..|cbn]. split; auto; ring. Qed. Instance poly_scalar_mult_dist_r: RightHeteroDistribute (·) (+) (+). Proof. intros a b x. induction x as [|x p IH]; [easy|cbn]. split; auto; ring. Qed. Instance poly_scalar_mult_assoc: HeteroAssociative (·) (·) (·) (.*.). Proof. intros a b x. induction x as [|p x IH]; [easy|cbn]. split; auto; ring. Qed. Lemma poly_scalar_mult_0_l q c: c = 0 → poly_eq_zero (c · q). Proof. intros ->. induction q as [|x q IH]; [easy|cbn]. rewrite poly_eq_zero_cons; split; auto. ring. Qed. Global Instance: Mult poly := fix F p q := match p with | [] => 0 | c :: p1 => c · q + (0 :: F p1 q) end. Lemma poly_mult_0_l p q: poly_eq_zero p → poly_eq_zero (p * q). Proof. induction 1 using poly_eq_zero_ind; [easy|cbn]. apply poly_eq_zero_plus. - now apply poly_scalar_mult_0_l. - rewrite poly_eq_zero_cons; auto. Qed. Lemma poly_mult_0_r p q: poly_eq_zero q → poly_eq_zero (p * q). Proof. induction p as [|x p IH]; [easy|cbn]. intro eq0. apply poly_eq_zero_plus. - now apply poly_scalar_mult_0_r. - rewrite poly_eq_zero_cons; auto. Qed. Instance poly_mult_proper: Proper ((=) ==> (=) ==> (=)) (.*.). Proof. refine (poly_eq_ind _ _ _). { intros ?? zp zp' q q' eqq. now apply poly_eq_zero_trans; apply poly_mult_0_l. } intros ???? eqx eqp IH q q' eqq. cbn. apply poly_plus_proper. { now rewrite eqx, eqq. } split; auto. now apply IH. Qed. Instance poly_mult_left_distr: LeftDistribute (.*.) (+). Proof. intros p q r. induction p as [|x p IH]; [easy|cbn]. rewrite (distribute_l x q r ). rewrite <- !associativity; apply poly_plus_proper; [easy|]. rewrite associativity, (commutativity (0::p*q)), <- associativity. apply poly_plus_proper; [easy|]. cbn; split; [ring|easy]. Qed. Lemma poly_mult_cons_r p q x: p * (x::q) = x · p + (0 :: p * q). Proof. induction p as [|y p IH]; cbn; auto. split; auto. rewrite IH, !associativity, (commutativity (y · q)). split; try easy. ring. Qed. Instance poly_mult_comm: Commutative (.*.). Proof. intros p. induction p as [|x p IH]. { cbn; intro. now apply poly_mult_0_r. } intros [|y q]; cbn. { rewrite poly_eq_zero_cons; split; auto. now apply poly_mult_0_r. } split. 1: ring. rewrite !poly_mult_cons_r, !associativity. apply poly_plus_proper. { apply commutativity. } rewrite <- poly_eq_cons; split; auto. apply IH. Qed. Instance poly_mult_right_distr: RightDistribute (.*.) (+). Proof. intros p q r. now rewrite commutativity, distribute_l, !(commutativity r). Qed. Instance poly_mult_1_l: LeftIdentity (.*.) 1. Proof. intro; cbn. rewrite poly_scalar_mult_1_l, poly_eq_zero_plus_r; auto. now split. Qed. Instance poly_mult_1_r: RightIdentity (.*.) 1. Proof. intro; rewrite commutativity; apply left_identity. Qed. Instance poly_mult_assoc: Associative (.*.). Proof with (try easy); cbn. intros x. induction x as [|x p IH]... intros q r; cbn. rewrite distribute_r. apply poly_plus_proper; cbn -[poly_eq]. { clear IH. induction q as [|y q IH]... rewrite distribute_l, <- associativity. apply poly_plus_proper... split; auto. ring. } assert (0 · r = 0) as ->. { now rewrite poly_eq_p_zero; apply poly_scalar_mult_0_l. } cbn; split; eauto. now rewrite IH. Qed. Global Instance poly_ring: Ring poly. Proof. repeat (split; try apply _). Qed. Global Instance poly_module : @Module R poly Ae Aplus Amult Azero Aone Anegate poly_eq (+) 0 (-) (·). Proof. split; apply _. Qed. End contents. (* Section test. Context `{Ring R} (x y: poly (poly (poly (poly R)))). Goal x + y == x * y. set (d := Plus_instance_0 ). set (u := Mult_instance_0). set (t := poly (poly R)). unfold poly_zero. *) math-classes-8.19.0/implementations/positive_semiring_elements.v000066400000000000000000000042161460576051100252370ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.additional_operations MathClasses.interfaces.orders MathClasses.interfaces.integers MathClasses.orders.semirings MathClasses.theory.shiftl MathClasses.theory.int_pow. Local Existing Instance pseudo_srorder_semiring. Section positive_semiring_elements. Context `{FullPseudoSemiRingOrder R} `{!PropHolds (1 ≶ 0)}. Add Ring R : (rings.stdlib_semiring_theory R). (* * Embedding of R₊ into R *) Global Instance Pos_inject: Cast (R₊) R := @proj1_sig R _. (* Operations *) Global Program Instance Pos_plus: Plus (R₊) := λ x y, (`x + `y)↾_. Next Obligation. destruct x as [x Hx], y as [y Hy]. now apply pos_plus_compat. Qed. Global Program Instance Pos_mult: Mult (R₊) := λ x y, (`x * `y)↾_. Next Obligation with auto. destruct x as [x Hx], y as [y Hy]. now apply pos_mult_compat. Qed. Global Program Instance Pos_1: One (R₊) := (1 : R). Next Obligation. now apply lt_0_1. Qed. (* * Equalitity *) Local Ltac unfold_equivs := unfold equiv, sig_equiv in *; simpl in *. Global Instance: Proper ((=) ==> (=) ==> (=)) Pos_plus. Proof. intros [x1 Ex1] [y1 Ey1] E1 [x2 Ex2] [y2 Ey2] E2. unfold_equivs. now rewrite E1, E2. Qed. Global Instance: Proper ((=) ==> (=) ==> (=)) Pos_mult. Proof. intros [x1 Ex1] [y1 Ey1] E1 [x2 Ex2] [y2 Ey2] E2. unfold_equivs. now rewrite E1, E2. Qed. Instance: Proper ((=) ==> (=)) Pos_inject. Proof. now repeat intro. Qed. Global Instance: Injective Pos_inject. Proof. now repeat (split; try apply _). Qed. Section shiftl. Context `{SemiRing B} `{!Biinduction B} `{!ShiftLSpec R B sl}. Global Program Instance Pos_shiftl: ShiftL (R₊) B | 5 := λ x n, (`x ≪ n)↾_. Next Obligation. destruct x. now apply shiftl_pos. Qed. End shiftl. End positive_semiring_elements. Section int_pow. Context `{DecField R} `{Apart R} `{!FullPseudoSemiRingOrder Rle Rlt} `{!TrivialApart R} `{∀ x y : R, Decision (x = y)} `{Integers B} `{!IntPowSpec R B ipw}. Global Program Instance Pos_int_pow: Pow (R₊) B | 5 := λ x n, (`x ^ n)↾_. Next Obligation. destruct x. now apply int_pow_pos. Qed. End int_pow. math-classes-8.19.0/implementations/semiring_pairs.v000066400000000000000000000265431460576051100226260ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.orders.rings. Inductive SRpair (SR : Type) := C { pos : SR ; neg : SR }. Arguments C {SR} _ _. Arguments pos {SR} _. Arguments neg {SR} _. Section semiring_pairs. Context `{SemiRing SR} `{Apart SR}. Context `{∀ z, LeftCancellation (+) z}. Add Ring SR : (rings.stdlib_semiring_theory SR). (* Equivalence *) Global Instance SRpair_equiv : Equiv (SRpair SR) | 4 := λ x y, pos x + neg y = pos y + neg x. Global Instance SRpair_apart `{Apart SR} : Apart (SRpair SR) := λ x y, pos x + neg y ≶ pos y + neg x. Global Instance SRpair_trivial_apart `{!TrivialApart SR} : TrivialApart (SRpair SR). Proof. intros x y. now Tactics.rapply trivial_apart. Qed. Instance: Setoid (SRpair SR). Proof. split; red; unfold equiv, SRpair_equiv. reflexivity. intros. now symmetry. intros x y z E E'. rewrite commutativity. rewrite (commutativity (pos z)). apply (left_cancellation (+) (pos y)). rewrite 2!associativity. rewrite <- E, E'. ring. Qed. Instance: Proper ((=) ==> (=) ==> (=)) C. Proof. intros x1 y1 E1 x2 y2 E2. unfold equiv, SRpair_equiv. simpl. now rewrite E1, E2. Qed. (* injection from SR *) Global Instance SRpair_inject: Cast SR (SRpair SR) := λ r, C r 0. Global Instance: Proper ((=) ==> (=)) SRpair_inject. Proof. intros x1 x2 E. unfold equiv, SRpair_equiv. simpl. now rewrite E. Qed. (* Relations, operations and constants *) Global Instance SRpair_plus: Plus (SRpair SR) := λ x y, C (pos x + pos y) (neg x + neg y). Global Instance SRpair_negate: Negate (SRpair SR) := λ x, C (neg x) (pos x). Global Instance SRpair_0: Zero (SRpair SR) := ('0 : SRpair SR). Global Instance SRpair_mult: Mult (SRpair SR) := λ x y, C (pos x * pos y + neg x * neg y) (pos x * neg y + neg x * pos y). Global Instance SRpair_1: One (SRpair SR) := ('1 : SRpair SR). Ltac unfolds := unfold SRpair_negate, SRpair_plus, equiv, SRpair_equiv in *; simpl in *. Ltac ring_on_sr := repeat intro; unfolds; try ring. Instance: Proper ((=) ==> (=)) SRpair_negate. Proof. intros x y E. unfolds. rewrite commutativity, <- E. ring. Qed. Instance: Proper ((=) ==> (=) ==> (=)) SRpair_plus. Proof with try ring. intros x1 y1 E1 x2 y2 E2. unfolds. transitivity (pos x1 + neg y1 + (pos x2 + neg y2))... rewrite E1, E2... Qed. Let SRpair_mult_proper_r (x y z : SRpair SR) : x = y → z * x = z * y. Proof with try ring. intros E. unfolds. transitivity (pos z * (pos x + neg y) + neg z * (pos y + neg x))... transitivity (pos z * (pos y + neg x) + neg z * (pos x + neg y))... now rewrite E. Qed. Instance: Commutative SRpair_mult. Proof. repeat intro. ring_on_sr. Qed. Instance: Proper ((=) ==> (=) ==> (=)) SRpair_mult. Proof. intros x1 y1 E1 x2 y2 E2. transitivity (x1 * y2). now apply SRpair_mult_proper_r. rewrite !(commutativity _ y2). now apply SRpair_mult_proper_r. Qed. Global Instance: Ring (SRpair SR). Proof. repeat (split; try apply _); ring_on_sr. Qed. (* A final word about inject *) Global Instance: SemiRing_Morphism SRpair_inject. Proof. repeat (constructor; try apply _); try reflexivity. intros x y. change (x + y + (0 + 0) = x + y + 0). ring. intros x y. change (x * y + (x * 0 + 0 * y) = x * y + 0 * 0 + 0). ring. Qed. Global Instance: Injective SRpair_inject. Proof. repeat (constructor; try apply _). intros x y. unfolds. now rewrite 2!rings.plus_0_r. Qed. Lemma SRpair_splits n m : C n m = 'n + -'m. Proof. ring_on_sr. Qed. Global Instance SRpair_le `{Le SR} : Le (SRpair SR) := λ x y, pos x + neg y ≤ pos y + neg x. Global Instance SRpair_lt `{Lt SR} : Lt (SRpair SR) := λ x y, pos x + neg y < pos y + neg x. Ltac unfold_le := unfold le, SRpair_le, equiv, SRpair_equiv; simpl. Ltac unfold_lt := unfold lt, SRpair_lt, equiv, SRpair_equiv; simpl. Section with_semiring_order. Context `{!SemiRingOrder SRle}. Instance: Proper ((=) ==> (=) ==> iff) SRpair_le. Proof. assert (∀ x1 y1 : SRpair SR, x1 = y1 → ∀ x2 y2, x2 = y2 → x1 ≤ x2 → y1 ≤ y2) as E. unfold_le. intros [xp1 xn1] [yp1 yn1] E1 [xp2 xn2] [yp2 yn2] E2 F. simpl in *. apply (order_reflecting (+ (xp2 + xn1))). setoid_replace (yp1 + yn2 + (xp2 + xn1)) with ((yp1 + xn1) + (xp2 + yn2)) by ring. rewrite <-E1, E2. setoid_replace (xp1 + yn1 + (yp2 + xn2)) with ((yp2 + yn1) + (xp1 + xn2)) by ring. now apply (order_preserving _). split; repeat intro; eapply E; eauto; symmetry; eauto. Qed. Instance: Reflexive SRpair_le. Proof. intros [? ?]. unfold_le. reflexivity. Qed. Instance: Transitive SRpair_le. Proof. intros [xp xn] [yp yn] [zp zn] E1 E2. unfold SRpair_le in *. simpl in *. apply (order_reflecting (+ (yn + yp))). setoid_replace (xp + zn + (yn + yp)) with ((xp + yn) + (yp + zn)) by ring. setoid_replace (zp + xn + (yn + yp)) with ((yp + xn) + (zp + yn)) by ring. now apply plus_le_compat. Qed. Instance: AntiSymmetric SRpair_le. Proof. intros [xp xn] [yp yn] E1 E2. unfold_le. now apply (antisymmetry (≤)). Qed. Instance: PartialOrder SRpair_le. Proof. repeat (split; try apply _). Qed. Global Instance: OrderEmbedding SRpair_inject. Proof. repeat (split; try apply _). intros x y E. unfold_le. simpl. now rewrite 2!rings.plus_0_r. intros x y E. unfold le, SRpair_le in E. simpl in E. now rewrite 2!rings.plus_0_r in E. Qed. Instance: ∀ z : SRpair SR, OrderPreserving ((+) z). Proof. repeat (split; try apply _). unfold_le. destruct z as [zp zn]. intros [xp xn] [yp yn] E. simpl in *. setoid_replace (zp + xp + (zn + yn)) with ((zp + zn) + (xp + yn)) by ring. setoid_replace (zp + yp + (zn + xn)) with ((zp + zn) + (yp + xn)) by ring. now apply (order_preserving _). Qed. Instance: ∀ x y : SRpair SR, PropHolds (0 ≤ x) → PropHolds (0 ≤ y) → PropHolds (0 ≤ x * y). Proof. intros [xp xn] [yp yn]. unfold PropHolds. unfold_le. intros E1 E2. ring_simplify in E1. ring_simplify in E2. destruct (decompose_le E1) as [a [Ea1 Ea2]], (decompose_le E2) as [b [Eb1 Eb2]]. rewrite Ea2, Eb2. ring_simplify. apply compose_le with (a * b). now apply nonneg_mult_compat. ring. Qed. Global Instance: SemiRingOrder SRpair_le. Proof. apply rings.from_ring_order; apply _. Qed. End with_semiring_order. Section with_strict_semiring_order. Context `{!StrictSemiRingOrder SRle}. Instance: Proper ((=) ==> (=) ==> iff) SRpair_lt. Proof. assert (∀ x1 y1 : SRpair SR, x1 = y1 → ∀ x2 y2, x2 = y2 → x1 < x2 → y1 < y2) as E. unfold_lt. intros [xp1 xn1] [yp1 yn1] E1 [xp2 xn2] [yp2 yn2] E2 F. simpl in *. apply (strictly_order_reflecting (+ (xp2 + xn1))). setoid_replace (yp1 + yn2 + (xp2 + xn1)) with ((yp1 + xn1) + (xp2 + yn2)) by ring. rewrite <-E1, E2. setoid_replace (xp1 + yn1 + (yp2 + xn2)) with ((yp2 + yn1) + (xp1 + xn2)) by ring. now apply (strictly_order_preserving _). split; repeat intro; eapply E; eauto; symmetry; eauto. Qed. Instance: Irreflexive SRpair_lt. Proof. intros [? ?] E. edestruct (irreflexivity (<)); eauto. Qed. Instance: Transitive SRpair_lt. Proof. intros [xp xn] [yp yn] [zp zn] E1 E2. unfold SRpair_lt in *. simpl in *. apply (strictly_order_reflecting (+ (yn + yp))). setoid_replace (xp + zn + (yn + yp)) with ((xp + yn) + (yp + zn)) by ring. setoid_replace (zp + xn + (yn + yp)) with ((yp + xn) + (zp + yn)) by ring. now apply plus_lt_compat. Qed. Instance: ∀ z : SRpair SR, StrictlyOrderPreserving ((+) z). Proof. repeat (split; try apply _). unfold_lt. destruct z as [zp zn]. intros [xp xn] [yp yn] E. simpl in *. setoid_replace (zp + xp + (zn + yn)) with ((zp + zn) + (xp + yn)) by ring. setoid_replace (zp + yp + (zn + xn)) with ((zp + zn) + (yp + xn)) by ring. now apply (strictly_order_preserving _). Qed. Instance: StrictSetoidOrder SRpair_lt. Proof. repeat (split; try apply _). Qed. Instance: ∀ x y : SRpair SR, PropHolds (0 < x) → PropHolds (0 < y) → PropHolds (0 < x * y). Proof. intros [xp xn] [yp yn]. unfold PropHolds. unfold_lt. intros E1 E2. ring_simplify in E1. ring_simplify in E2. destruct (decompose_lt E1) as [a [Ea1 Ea2]], (decompose_lt E2) as [b [Eb1 Eb2]]. rewrite Ea2, Eb2. ring_simplify. apply compose_lt with (a * b). now apply pos_mult_compat. ring. Qed. Global Instance: StrictSemiRingOrder SRpair_lt. Proof. apply from_strict_ring_order; apply _. Qed. End with_strict_semiring_order. Section with_full_pseudo_semiring_order. Context `{!FullPseudoSemiRingOrder SRle SRlt}. Instance: StrongSetoid SR := pseudo_order_setoid. Instance: StrongSetoid (SRpair SR). Proof. split. intros [??] E. now eapply (irreflexivity (≶)); eauto. intros [??] [??] E. unfold apart, SRpair_apart. now symmetry. intros [xp xn] [yp yn] E [zp zn]. unfold apart, SRpair_apart in *. simpl in *. apply (strong_left_cancellation (+) zn) in E. edestruct (cotransitive E). left. apply (strong_extensionality (+ yn)). setoid_replace (xp + zn + yn) with (zn + (xp + yn)) by ring. eassumption. right. apply (strong_extensionality (+ xn)). setoid_replace (zp + yn + xn) with (zp + xn + yn) by ring. setoid_replace (yp + zn + xn) with (zn + (yp + xn)) by ring. eassumption. intros [??] [??]. now Tactics.rapply tight_apart. Qed. Instance: FullPseudoOrder SRpair_le SRpair_lt. Proof. split. split; try apply _. intros [??] [??]. unfold_lt. now apply pseudo_order_antisym. intros [xp xn] [yp yn] E [zp zn]. unfold lt, SRpair_lt in *. simpl in *. apply (strictly_order_preserving (zn +)) in E. edestruct (cotransitive E). left. apply (strictly_order_reflecting (+ yn)). setoid_replace (xp + zn + yn) with (zn + (xp + yn)) by ring. eassumption. right. apply (strictly_order_reflecting (+ xn)). setoid_replace (zp + yn + xn) with (zp + xn + yn) by ring. setoid_replace (yp + zn + xn) with (zn + (yp + xn)) by ring. eassumption. intros [??] [??]. now Tactics.rapply apart_iff_total_lt. intros [??] [??]. now Tactics.rapply le_iff_not_lt_flip. Qed. Instance: ∀ z : SRpair SR, StrongSetoid_Morphism (z *.). Proof. intros [zp zn]. split; try apply _. intros [xp xn] [yp yn] E1. unfold apart, SRpair_apart in *. simpl in *. destruct (strong_binary_extensionality (+) (zp * (xp + yn)) (zn * (yp + xn)) (zp * (yp + xn)) (zn * (xp + yn))). eapply strong_setoids.apart_proper; eauto; ring. now apply (strong_extensionality (zp *.)). symmetry. now apply (strong_extensionality (zn *.)). Qed. Global Instance: FullPseudoSemiRingOrder SRpair_le SRpair_lt. Proof. apply from_full_pseudo_ring_order; try apply _. now apply strong_setoids.strong_binary_setoid_morphism_commutative. Qed. End with_full_pseudo_semiring_order. Global Instance SRpair_dec `{∀ x y : SR, Decision (x = y)} : ∀ x y : SRpair SR, Decision (x = y) := λ x y, decide_rel (=) (pos x + neg y) (pos y + neg x). Global Program Instance SRpair_le_dec `{Le SR} `{∀ x y: SR, Decision (x ≤ y)} : ∀ x y : SRpair SR, Decision (x ≤ y) := λ x y, match decide_rel (≤) (pos x + neg y) (pos y + neg x) with | left E => left _ | right E => right _ end. End semiring_pairs. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Opaque SRpair_equiv. #[global] Typeclasses Opaque SRpair_le. math-classes-8.19.0/implementations/stdlib_binary_integers.v000066400000000000000000000154431460576051100243350ustar00rootroot00000000000000Require MathClasses.interfaces.naturals MathClasses.theory.naturals MathClasses.implementations.peano_naturals MathClasses.theory.integers. Require Import Coq.ZArith.BinInt Coq.setoid_ring.Ring Coq.Arith.Arith Coq.NArith.NArith Coq.ZArith.ZArith Coq.Numbers.Integer.Binary.ZBinary MathClasses.interfaces.abstract_algebra MathClasses.interfaces.integers MathClasses.implementations.natpair_integers MathClasses.implementations.stdlib_binary_naturals MathClasses.interfaces.additional_operations MathClasses.interfaces.orders MathClasses.implementations.nonneg_integers_naturals. (* canonical names: *) #[global] Instance Z_equiv: Equiv Z := eq. #[global] Instance Z_plus: Plus Z := Zplus. #[global] Instance Z_0: Zero Z := 0%Z. #[global] Instance Z_1: One Z := 1%Z. #[global] Instance Z_mult: Mult Z := Zmult. #[global] Instance Z_negate: Negate Z := Z.opp. (* some day we'd like to do this with [Existing Instance] *) #[global] Instance: Ring Z. Proof. repeat (split; try apply _); repeat intro. now apply Zplus_assoc. now apply Zplus_0_r. now apply Zplus_opp_l. now apply Zplus_opp_r. now apply Zplus_comm. now apply Zmult_assoc. now apply Zmult_1_l. now apply Zmult_1_r. now apply Zmult_comm. now apply Zmult_plus_distr_r. Qed. (* misc: *) #[global] Instance: ∀ x y : Z, Decision (x = y) := Z.eq_dec. Add Ring Z: (rings.stdlib_ring_theory Z). (* * Embedding N into Z *) #[global] Instance inject_N_Z: Cast N Z := Z_of_N. #[global] Instance: SemiRing_Morphism Z_of_N. Proof. repeat (split; try apply _). exact Znat.Z_of_N_plus. exact Znat.Z_of_N_mult. Qed. #[global] Instance: Injective Z_of_N. Proof. repeat (split; try apply _). intros x y E. now apply Znat.Z_of_N_eq_iff. Qed. (* SRpair N and Z are isomorphic *) Definition Npair_to_Z (x : SRpair N) : Z := ('pos x - 'neg x)%mc. #[global] Instance: Proper (=) Npair_to_Z. Proof. intros [xp xn] [yp yn] E; do 2 red in E; unfold Npair_to_Z; simpl in *. apply (right_cancellation (+) ('yn + 'xn)); ring_simplify. now rewrite <-?rings.preserves_plus, E, commutativity. Qed. #[global] Instance: SemiRing_Morphism Npair_to_Z. Proof. repeat (split; try apply _). intros [xp xn] [yp yn]. change ('(xp + yp) - '(xn + yn) = 'xp - 'xn + ('yp - 'yn)). rewrite ?rings.preserves_plus. ring. intros [xp xn] [yp yn]. change ('(xp * yp + xn * yn) - '(xp * yn + xn * yp) = ('xp - 'xn) * ('yp - 'yn)). rewrite ?rings.preserves_plus, ?rings.preserves_mult. ring. Qed. #[global] Instance: Injective Npair_to_Z. Proof. split; try apply _. intros [xp xn] [yp yn] E. unfold Npair_to_Z in E. do 2 red. simpl in *. apply (injective (cast N Z)). rewrite ?rings.preserves_plus. apply (right_cancellation (+) ('xp - 'xn)). rewrite E at 1. ring. Qed. #[global] Instance Z_to_Npair: Inverse Npair_to_Z := λ x, match x with | Z0 => C 0 0 | Zpos p => C (Npos p) 0 | Zneg p => C 0 (Npos p) end. #[global] Instance: Surjective Npair_to_Z. Proof. split; try apply _. intros [|?|?] ? E; now rewrite <-E. Qed. #[global] Instance: Bijective Npair_to_Z := {}. #[global] Instance: SemiRing_Morphism Z_to_Npair. Proof. change (SemiRing_Morphism (Npair_to_Z⁻¹)). split; apply _. Qed. #[global] Instance: IntegersToRing Z := integers.retract_is_int_to_ring Npair_to_Z. #[global] Instance: Integers Z := integers.retract_is_int Npair_to_Z. #[global] Instance Z_le: Le Z := Z.le. #[global] Instance Z_lt: Lt Z := Z.lt. #[global] Instance: SemiRingOrder Z_le. Proof. assert (PartialOrder Z_le). repeat (split; try apply _). exact Zorder.Zle_antisym. rapply rings.from_ring_order. repeat (split; try apply _). intros x y E. now apply Zorder.Zplus_le_compat_l. intros x E y F. now apply Zorder.Zmult_le_0_compat. Qed. #[global] Instance: TotalRelation Z_le. Proof. intros x y. destruct (Zorder.Zle_or_lt x y); intuition. right. now apply Zorder.Zlt_le_weak. Qed. #[global] Instance: FullPseudoSemiRingOrder Z_le Z_lt. Proof. rapply semirings.dec_full_pseudo_srorder. split. intro. split. now apply Zorder.Zlt_le_weak. now apply Zorder.Zlt_not_eq. intros [E1 E2]. destruct (Zorder.Zle_lt_or_eq _ _ E1). easy. now destruct E2. Qed. (* * Embedding of the Peano naturals into [Z] *) #[global] Instance inject_nat_Z: Cast nat Z := Z_of_nat. #[global] Instance: SemiRing_Morphism Z_of_nat. Proof. repeat (split; try apply _). exact Znat.inj_plus. exact Znat.inj_mult. Qed. (* absolute value *) #[global] Program Instance Z_abs_nat: IntAbs Z nat := λ x, match x with | Z0 => inl (0:nat) | Zpos p => inl (nat_of_P p) | Zneg p => inr (nat_of_P p) end. Next Obligation. reflexivity. Qed. Next Obligation. now rewrite <-(naturals.to_semiring_unique Z_of_nat), Znat.Z_of_nat_of_P. Qed. Next Obligation. now rewrite <-(naturals.to_semiring_unique Z_of_nat), Znat.Z_of_nat_of_P. Qed. #[global] Program Instance Z_abs_N: IntAbs Z N := λ x, match x with | Z0 => inl (0:N) | Zpos p => inl (Npos p) | Zneg p => inr (Npos p) end. Next Obligation. reflexivity. Qed. Next Obligation. now rewrite <-(naturals.to_semiring_unique Z_of_N). Qed. Next Obligation. now rewrite <-(naturals.to_semiring_unique Z_of_N). Qed. (* Efficient nat_pow *) #[global] Program Instance Z_pow: Pow Z (Z⁺) := Z.pow. #[global] Instance: NatPowSpec Z (Z⁺) Z_pow. Proof. split; unfold pow, Z_pow. intros x1 y1 E1 [x2 Ex2] [y2 Ey2] E2. unfold equiv, sig_equiv in E2. simpl in *. now rewrite E1, E2. intros. now apply Z.pow_0_r. intros x n. rewrite rings.preserves_plus, rings.preserves_1. rewrite <-(Z.pow_1_r x) at 2. apply Z.pow_add_r. auto with zarith. now destruct n. Qed. #[global] Instance Z_Npow: Pow Z N := λ x n, Z.pow x ('n). #[global] Instance: NatPowSpec Z N Z_Npow. Proof. split; unfold pow, Z_Npow. solve_proper. intros. now apply Z.pow_0_r. intros x n. rewrite rings.preserves_plus, rings.preserves_1. rewrite <-(Z.pow_1_r x) at 2. apply Z.pow_add_r. auto with zarith. now destruct n. Qed. (* Efficient shiftl *) #[global] Program Instance Z_shiftl: ShiftL Z (Z⁺) := Z.shiftl. #[global] Instance: ShiftLSpec Z (Z⁺) Z_shiftl. Proof. apply shiftl_spec_from_nat_pow. intros x [n En]. apply Z.shiftl_mul_pow2. now apply En. Qed. #[global] Instance Z_Nshiftl: ShiftL Z N := λ x n, Z.shiftl x ('n). #[global] Instance: ShiftLSpec Z N Z_Nshiftl. Proof. apply shiftl_spec_from_nat_pow. intros x n. apply Z.shiftl_mul_pow2. now destruct n. Qed. #[global] Program Instance Z_abs: Abs Z := Z.abs. Next Obligation. split; intros E. now apply Z.abs_eq. now apply Z.abs_neq. Qed. #[global] Instance Z_div: DivEuclid Z := Z.div. #[global] Instance Z_mod: ModEuclid Z := Zmod. #[global] Instance: EuclidSpec Z _ _. Proof. split; try apply _. exact Z.div_mod. intros x y Ey. destruct (Z_mod_remainder x y); intuition. now intros []. now intros []. Qed. math-classes-8.19.0/implementations/stdlib_binary_naturals.v000066400000000000000000000057471460576051100243540ustar00rootroot00000000000000Require Import Coq.NArith.NArith MathClasses.implementations.peano_naturals MathClasses.theory.naturals MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals MathClasses.interfaces.orders MathClasses.interfaces.additional_operations. (* canonical names for relations/operations/constants: *) #[global] Instance N_equiv : Equiv N := eq. #[global] Instance N_0 : Zero N := 0%N. #[global] Instance N_1 : One N := 1%N. #[global] Instance N_plus : Plus N := Nplus. #[global] Instance N_mult : Mult N := Nmult. (* properties: *) #[global] Instance: SemiRing N. Proof. repeat (split; try apply _); repeat intro. now apply Nplus_assoc. now apply Nplus_0_r. now apply Nplus_comm. now apply Nmult_assoc. now apply Nmult_1_l. now apply Nmult_1_r. now apply Nmult_comm. now apply Nmult_plus_distr_l. Qed. #[global] Instance: ∀ x y : N, Decision (x = y) := N.eq_dec. #[global] Instance inject_nat_N: Cast nat N := N_of_nat. #[global] Instance inject_N_nat: Cast N nat := nat_of_N. #[global] Instance: SemiRing_Morphism nat_of_N. Proof. repeat (split; try apply _); repeat intro. now apply nat_of_Nplus. now apply nat_of_Nmult. Qed. #[global] Instance: Inverse nat_of_N := N_of_nat. #[global] Instance: Surjective nat_of_N. Proof. constructor. intros x y E. rewrite <- E. now apply nat_of_N_of_nat. now apply _. Qed. #[global] Instance: Injective nat_of_N. Proof. constructor. exact nat_of_N_inj. apply _. Qed. #[global] Instance: Bijective nat_of_N := {}. #[global] Instance: Inverse N_of_nat := nat_of_N. #[global] Instance: Bijective N_of_nat. Proof. apply jections.flip_bijection. Qed. #[global] Instance: SemiRing_Morphism N_of_nat. Proof. change (SemiRing_Morphism (nat_of_N⁻¹)). split; apply _. Qed. #[global] Instance: NaturalsToSemiRing N := retract_is_nat_to_sr N_of_nat. #[global] Instance: Naturals N := retract_is_nat N_of_nat. (* order *) #[global] Instance N_le: Le N := N.le. #[global] Instance N_lt: Lt N := N.lt. #[global] Instance: FullPseudoSemiRingOrder N_le N_lt. Proof. assert (PartialOrder N_le). repeat (split; try apply _). exact N.le_antisymm. assert (SemiRingOrder N_le). split; try apply _. intros x y E. exists (Nminus y x). symmetry. rewrite commutativity. now apply N.sub_add. repeat (split; try apply _); intros. now apply N.add_le_mono_l. eapply N.add_le_mono_l. eassumption. intros. now apply Nle_0. assert (TotalRelation N_le). intros x y. now apply N.le_ge_cases. rapply semirings.dec_full_pseudo_srorder. split. intro. now apply N.le_neq. intros [E1 E2]. now apply N.Private_Tac.le_neq_lt. Qed. #[global] Program Instance: ∀ x y: N, Decision (x ≤ y) := λ y x, match N.compare y x with | Gt => right _ | _ => left _ end. Next Obligation. now apply not_symmetry. Qed. #[global] Instance N_cut_minus: CutMinus N := Nminus. #[global] Instance: CutMinusSpec N _. Proof. split; try apply _. intros. now apply N.sub_add. intros. now apply Nminus_N0_Nle. Qed. math-classes-8.19.0/implementations/stdlib_rationals.v000066400000000000000000000112671460576051100231450ustar00rootroot00000000000000Require MathClasses.implementations.stdlib_binary_integers Coq.setoid_ring.Field Coq.QArith.Qfield MathClasses.theory.rationals. Require Import Coq.ZArith.ZArith Coq.setoid_ring.Ring Coq.QArith.QArith_base Coq.QArith.Qabs Coq.QArith.Qpower MathClasses.interfaces.abstract_algebra MathClasses.interfaces.rationals MathClasses.interfaces.orders MathClasses.interfaces.additional_operations MathClasses.implementations.field_of_fractions MathClasses.orders.integers. (* canonical names for relations/operations/constants: *) #[global] Instance Q_eq: Equiv Q := Qeq. #[global] Instance Q_0 : Zero Q := 0%Q. #[global] Instance Q_1 : One Q := 1%Q. #[global] Instance Q_opp : Negate Q := Qopp. #[global] Instance Q_plus : Plus Q := Qplus. #[global] Instance Q_mult : Mult Q := Qmult. #[global] Instance Q_recip : DecRecip Q := Qinv. (* properties: *) #[global] Instance: Setoid Q := {}. #[global] Instance: DecField Q. Proof dec_fields.from_stdlib_field_theory Qfield.Qsft eq_refl. (* misc: *) #[global] Instance: ∀ x y: Q, Decision (x = y) := Qeq_dec. #[global] Instance inject_Z_Q: Cast Z Q := inject_Z. #[global] Instance: Proper ((=) ==> (=)) inject_Z. Proof. intros x y H. unfold inject_Z. repeat red. simpl. now rewrite H. Qed. #[global] Instance: SemiRing_Morphism inject_Z. Proof. repeat (split; try apply _). intros x y. repeat red. simpl. now rewrite ?Zmult_1_r. Qed. #[global] Instance: Injective inject_Z. Proof. constructor. 2: apply _. intros x y. change (x * 1 = y * 1 → x = y). rewrite 2!rings.mult_1_r. intuition. Qed. #[global] Program Instance Q_to_fracZ: Cast Q (Frac Z) := λ x, frac (Qnum x) (Zpos (Qden x)) _. #[global] Instance: Proper ((=) ==> (=)) Q_to_fracZ. Proof. intros ? ? E. easy. Qed. #[global] Instance: SemiRing_Morphism Q_to_fracZ. Proof. repeat (split; try apply _). Qed. #[global] Instance: Injective Q_to_fracZ. Proof. split; try apply _. intros ? ? E. easy. Qed. #[global] Instance: RationalsToFrac Q := rationals.alt_to_frac Q_to_fracZ. #[global] Instance: Rationals Q := rationals.alt_Build_Rationals Q_to_fracZ inject_Z. (* order: *) #[global] Instance Q_le: Le Q := Qle. #[global] Instance Q_lt: Lt Q := Qlt. #[global] Instance: SemiRingOrder Q_le. Proof. assert (PartialOrder Q_le). repeat (split; try apply _). exact Qle_refl. exact Qle_trans. exact Qle_antisym. apply (rings.from_ring_order (Rle:=Q_le)). repeat (split; try apply _). intros. apply Qplus_le_compat. now apply Qle_refl. easy. intros. now apply Qmult_le_0_compat. Qed. #[global] Instance: TotalRelation Q_le. Proof. intros x y. destruct (Qlt_le_dec x y); auto. left. now apply Qlt_le_weak. Qed. #[global] Instance: FullPseudoSemiRingOrder Q_le Q_lt. Proof. rapply (semirings.dec_full_pseudo_srorder (R:=Q)). split. intro. split. now apply Zorder.Zlt_le_weak. now apply Zorder.Zlt_not_eq. intros [E1 E2]. destruct (Zorder.Zle_lt_or_eq _ _ E1). easy. now destruct E2. Qed. #[global] Program Instance: ∀ x y: Q, Decision (x ≤ y) := λ y x, match Qlt_le_dec x y with | left _ => right _ | right _ => left _ end. Next Obligation. now apply Qlt_not_le. Qed. (* additional operations *) #[global] Program Instance: Abs Q := Qabs. Next Obligation. split; intros E. now apply Qabs_pos. now apply Qabs_neg. Qed. #[global] Instance Q_pow: Pow Q Z := Qpower. #[global] Instance: IntPowSpec Q Z Q_pow. Proof. split. apply _. reflexivity. exact Qpower_0. intros. now apply Qpower_plus. Qed. #[global] Instance Q_Npow: Pow Q N := λ x n, Qpower x (cast N Z n). #[global] Instance: NatPowSpec Q N Q_Npow. Proof. split. unfold pow, Q_Npow. solve_proper. reflexivity. intros. unfold pow, Q_Npow. rewrite rings.preserves_plus. rewrite Qpower_plus'. reflexivity. change (1 + cast N Z n ≠ 0). apply orders.lt_ne_flip. apply nat_int.le_iff_lt_S. apply nat_int.to_semiring_nonneg. Qed. #[global] Instance Q_shiftl: ShiftL Q Z := λ x k, match k with | Z0 => x | Zpos p => Qmake (Z.shiftl (Qnum x) (Zpos p)) (Qden x) | Zneg p => Qmake (Qnum x) (shift_pos p (Qden x)) end. #[global] Instance: ShiftLSpec Q Z Q_shiftl. Proof. apply shiftl_spec_from_int_pow. unfold shiftl, Q_shiftl. intros [n d] [|p|p]. change ((n#d) = (n#d) * 1). now rewrite rings.mult_1_r. unfold Qnum, Qden. rewrite !Qmake_Qdiv. unfold Qdiv. rewrite Z.shiftl_mul_pow2 by auto with zarith. rewrite Zmult_comm, inject_Z_mult, Zpower_Qpower by now destruct p. now rewrite <-Qmult_assoc, Qmult_comm. unfold Qnum, Qden. rewrite !Qmake_Qdiv. unfold Qdiv. rewrite shift_pos_correct. change (Zpower_pos 2 p) with (2 ^ (Zpos p))%Z. rewrite Zmult_comm, inject_Z_mult, Zpower_Qpower by now destruct p. now rewrite Qinv_mult_distr, Qmult_assoc. Qed. math-classes-8.19.0/interfaces/000077500000000000000000000000001460576051100163255ustar00rootroot00000000000000math-classes-8.19.0/interfaces/abstract_algebra.v000066400000000000000000000264201460576051100220000ustar00rootroot00000000000000Require Export MathClasses.interfaces.canonical_names MathClasses.misc.util MathClasses.misc.decision MathClasses.misc.propholds MathClasses.misc.workarounds MathClasses.misc.setoid_tactics. (* For various structures we omit declaration of substructures. For example, if we say: Class Setoid_Morphism := { setoidmor_a :: Setoid A ; setoidmor_b :: Setoid B ; sm_proper :: Proper ((=) ==> (=)) f }. then each time a Setoid instance is required, Coq will try to prove that a Setoid_Morphism exists. This obviously results in an enormous blow-up of the search space. Moreover, one should be careful to declare a Setoid_Morphisms as a substructure. Consider [f t1 t2], now if we want to perform setoid rewriting in [t2] Coq will first attempt to prove that [f t1] is Proper, for which it will attempt to prove [Setoid_Morphism (f t1)]. If many structures declare Setoid_Morphism as a substructure, setoid rewriting will become horribly slow. *) (* An unbundled variant of the former CoRN RSetoid *) Class Setoid A {Ae : Equiv A} : Prop := setoid_eq :: Equivalence (@equiv A Ae). (* An unbundled variant of the former CoRN CSetoid. We do not include a proof that A is a Setoid because it can be derived. *) Class StrongSetoid A {Ae : Equiv A} `{Aap : Apart A} : Prop := { strong_setoid_irreflexive :: Irreflexive (≶) ; strong_setoid_symmetric :: Symmetric (≶) ; strong_setoid_cotrans :: CoTransitive (≶) ; tight_apart : ∀ x y, ¬x ≶ y ↔ x = y }. Arguments tight_apart {A Ae Aap StrongSetoid} _ _. Section setoid_morphisms. Context {A B} {Ae : Equiv A} {Aap : Apart A} {Be : Equiv B} {Bap : Apart B} (f : A → B). Class Setoid_Morphism := { setoidmor_a : Setoid A ; setoidmor_b : Setoid B ; sm_proper :: Proper ((=) ==> (=)) f }. Class StrongSetoid_Morphism := { strong_setoidmor_a : StrongSetoid A ; strong_setoidmor_b : StrongSetoid B ; strong_extensionality : ∀ x y, f x ≶ f y → x ≶ y }. End setoid_morphisms. Arguments sm_proper {A B Ae Be f Setoid_Morphism} _ _ _. #[global] Hint Extern 4 (?f _ = ?f _) => unshelve eapply (sm_proper (f:=f)). Section setoid_binary_morphisms. Context {A B C} {Ae: Equiv A} {Aap: Apart A} {Be: Equiv B} {Bap : Apart B} {Ce: Equiv C} {Cap : Apart C} (f : A → B → C). Class StrongSetoid_BinaryMorphism := { strong_binary_setoidmor_a : StrongSetoid A ; strong_binary_setoidmor_b : StrongSetoid B ; strong_binary_setoidmor_c : StrongSetoid C ; strong_binary_extensionality : ∀ x₁ y₁ x₂ y₂, f x₁ y₁ ≶ f x₂ y₂ → x₁ ≶ x₂ ∨ y₁ ≶ y₂ }. End setoid_binary_morphisms. (* Since apartness usually only becomes relevant when considering fields (e.g. the real numbers), we do not include it in the lower part of the algebraic hierarchy (as opposed to CoRN). *) Section upper_classes. Context A {Ae : Equiv A}. Class SemiGroup {Aop: SgOp A} : Prop := { sg_setoid :: Setoid A ; sg_ass :: Associative (&) ; sg_op_proper :: Proper ((=) ==> (=) ==> (=)) (&) }. Class CommutativeSemiGroup {Aop : SgOp A} : Prop := { comsg_setoid :: @SemiGroup Aop ; comsg_ass :: Commutative (&) }. Class SemiLattice {Aop : SgOp A} : Prop := { semilattice_sg :: @CommutativeSemiGroup Aop ; semilattice_idempotent :: BinaryIdempotent (&)}. Class Monoid {Aop : SgOp A} {Aunit : MonUnit A} : Prop := { monoid_semigroup :: SemiGroup ; monoid_left_id :: LeftIdentity (&) mon_unit ; monoid_right_id :: RightIdentity (&) mon_unit }. Class CommutativeMonoid {Aop Aunit} : Prop := { commonoid_mon :: @Monoid Aop Aunit ; commonoid_commutative :: Commutative (&) }. Class Group {Aop Aunit} {Anegate : Negate A} : Prop := { group_monoid :: @Monoid Aop Aunit ; negate_proper :: Setoid_Morphism (-) ; negate_l :: LeftInverse (&) (-) mon_unit ; negate_r :: RightInverse (&) (-) mon_unit }. Class BoundedSemiLattice {Aop Aunit} : Prop := { bounded_semilattice_mon :: @CommutativeMonoid Aop Aunit ; bounded_semilattice_idempotent :: BinaryIdempotent (&)}. Class AbGroup {Aop Aunit Anegate} : Prop := { abgroup_group :: @Group Aop Aunit Anegate ; abgroup_commutative :: Commutative (&) }. Context {Aplus : Plus A} {Amult : Mult A} {Azero : Zero A} {Aone : One A}. Class SemiRing : Prop := { semiplus_monoid :: @CommutativeMonoid plus_is_sg_op zero_is_mon_unit ; semimult_monoid :: @CommutativeMonoid mult_is_sg_op one_is_mon_unit ; semiring_distr :: LeftDistribute (.*.) (+) ; semiring_left_absorb :: LeftAbsorb (.*.) 0 }. Context {Anegate : Negate A}. Class Ring : Prop := { ring_group :: @AbGroup plus_is_sg_op zero_is_mon_unit _ ; ring_monoid :: @CommutativeMonoid mult_is_sg_op one_is_mon_unit ; ring_dist :: LeftDistribute (.*.) (+) }. (* For now, we follow CoRN/ring_theory's example in having Ring and SemiRing require commutative multiplication. *) Class IntegralDomain : Prop := { intdom_ring : Ring ; intdom_nontrivial : PropHolds (1 ≠ 0) ; intdom_nozeroes :: NoZeroDivisors A }. (* We do not include strong extensionality for (-) and (/) because it can de derived *) Class Field {Aap: Apart A} {Arecip: Recip A} : Prop := { field_ring :: Ring ; field_strongsetoid :: StrongSetoid A ; field_plus_ext :: StrongSetoid_BinaryMorphism (+) ; field_mult_ext :: StrongSetoid_BinaryMorphism (.*.) ; field_nontrivial : PropHolds (1 ≶ 0) ; recip_proper :: Setoid_Morphism (//) ; recip_inverse : ∀ x, `x // x = 1 }. (* We let /0 = 0 so properties as Injective (/), f (/x) = / (f x), / /x = x, /x * /y = /(x * y) hold without any additional assumptions *) Class DecField {Adec_recip : DecRecip A} : Prop := { decfield_ring :: Ring ; decfield_nontrivial : PropHolds (1 ≠ 0) ; dec_recip_proper :: Setoid_Morphism (/) ; dec_recip_0 : /0 = 0 ; dec_recip_inverse : ∀ x, x ≠ 0 → x / x = 1 }. End upper_classes. (* Due to bug #2528 *) #[global] Hint Extern 4 (PropHolds (1 ≠ 0)) => eapply @intdom_nontrivial : typeclass_instances. #[global] Hint Extern 5 (PropHolds (1 ≶ 0)) => eapply @field_nontrivial : typeclass_instances. #[global] Hint Extern 5 (PropHolds (1 ≠ 0)) => eapply @decfield_nontrivial : typeclass_instances. (* For a strange reason Ring instances of Integers are sometimes obtained by Integers -> IntegralDomain -> Ring and sometimes directly. Making this an instance with a low priority instead of using intdom_ring:: Ring forces Coq to take the right way *) #[global] Hint Extern 10 (Ring _) => apply @intdom_ring : typeclass_instances. Arguments recip_inverse {A Ae Aplus Amult Azero Aone Anegate Aap Arecip Field} _. Arguments dec_recip_inverse {A Ae Aplus Amult Azero Aone Anegate Adec_recip DecField} _ _. Arguments dec_recip_0 {A Ae Aplus Amult Azero Aone Anegate Adec_recip DecField}. Arguments sg_op_proper {A Ae Aop SemiGroup} _ _ _ _ _ _. Section lattice. Context A `{Ae: Equiv A} {Ajoin: Join A} {Ameet: Meet A} `{Abottom : Bottom A}. Class JoinSemiLattice : Prop := join_semilattice :: @SemiLattice A Ae join_is_sg_op. Class BoundedJoinSemiLattice : Prop := bounded_join_semilattice :: @BoundedSemiLattice A Ae join_is_sg_op bottom_is_mon_unit. Class MeetSemiLattice : Prop := meet_semilattice :: @SemiLattice A Ae meet_is_sg_op. Class Lattice : Prop := { lattice_join :: JoinSemiLattice ; lattice_meet :: MeetSemiLattice ; join_meet_absorption :: Absorption (⊔) (⊓) ; meet_join_absorption :: Absorption (⊓) (⊔)}. Class DistributiveLattice : Prop := { distr_lattice_lattice :: Lattice ; join_meet_distr_l :: LeftDistribute (⊔) (⊓) }. End lattice. Class Category O `{!Arrows O} `{∀ x y: O, Equiv (x ⟶ y)} `{!CatId O} `{!CatComp O}: Prop := { arrow_equiv :: ∀ x y, Setoid (x ⟶ y) ; comp_proper :: ∀ x y z, Proper ((=) ==> (=) ==> (=)) (comp x y z) ; comp_assoc :: ArrowsAssociative O ; id_l :: ∀ x y, LeftIdentity (comp x y y) cat_id ; id_r :: ∀ x y, RightIdentity (comp x x y) cat_id }. (* note: no equality on objects. *) (* todo: use my comp everywhere *) Arguments comp_assoc {O arrows eq CatId CatComp Category w x y z} _ _ _ : rename. Section morphism_classes. Context {A B} {Ae : Equiv A} {Be : Equiv B}. Class SemiGroup_Morphism {Aop Bop} (f : A → B) := { sgmor_a : @SemiGroup A Ae Aop ; sgmor_b : @SemiGroup B Be Bop ; sgmor_setmor :: Setoid_Morphism f ; preserves_sg_op : ∀ x y, f (x & y) = f x & f y }. Class JoinSemiLattice_Morphism {Ajoin Bjoin} (f : A → B) := { join_slmor_a : @JoinSemiLattice A Ae Ajoin ; join_slmor_b : @JoinSemiLattice B Be Bjoin ; join_slmor_sgmor :: @SemiGroup_Morphism join_is_sg_op join_is_sg_op f }. Class MeetSemiLattice_Morphism {Ameet Bmeet} (f : A → B) := { meet_slmor_a : @MeetSemiLattice A Ae Ameet ; meet_slmor_b : @MeetSemiLattice B Be Bmeet ; meet_slmor_sgmor :: @SemiGroup_Morphism meet_is_sg_op meet_is_sg_op f }. Class Monoid_Morphism {Aunit Bunit Aop Bop} (f : A → B) := { monmor_a : @Monoid A Ae Aop Aunit ; monmor_b : @Monoid B Be Bop Bunit ; monmor_sgmor :: SemiGroup_Morphism f ; preserves_mon_unit : f mon_unit = mon_unit }. Class BoundedJoinSemiLattice_Morphism {Abottom Bbottom Ajoin Bjoin} (f : A → B) := { bounded_join_slmor_a : @BoundedJoinSemiLattice A Ae Ajoin Abottom ; bounded_join_slmor_b : @BoundedJoinSemiLattice B Be Bjoin Bbottom ; bounded_join_slmor_monmor :: @Monoid_Morphism bottom_is_mon_unit bottom_is_mon_unit join_is_sg_op join_is_sg_op f }. Class SemiRing_Morphism {Aplus Amult Azero Aone Bplus Bmult Bzero Bone} (f : A → B) := { semiringmor_a : @SemiRing A Ae Aplus Amult Azero Aone ; semiringmor_b : @SemiRing B Be Bplus Bmult Bzero Bone ; semiringmor_plus_mor :: @Monoid_Morphism zero_is_mon_unit zero_is_mon_unit plus_is_sg_op plus_is_sg_op f ; semiringmor_mult_mor :: @Monoid_Morphism one_is_mon_unit one_is_mon_unit mult_is_sg_op mult_is_sg_op f }. Class Lattice_Morphism {Ajoin Ameet Bjoin Bmeet} (f : A → B) := { latticemor_a : @Lattice A Ae Ajoin Ameet ; latticemor_b : @Lattice B Be Bjoin Bmeet ; latticemor_join_mor :: JoinSemiLattice_Morphism f ; latticemor_meet_mor :: MeetSemiLattice_Morphism f }. Context {Aap : Apart A} {Bap : Apart B}. Class StrongSemiRing_Morphism {Aplus Amult Azero Aone Bplus Bmult Bzero Bone} (f : A → B) := { strong_semiringmor_sr_mor :: @SemiRing_Morphism Aplus Amult Azero Aone Bplus Bmult Bzero Bone f ; strong_semiringmor_strong_mor :: StrongSetoid_Morphism f }. End morphism_classes. Section jections. Context {A B} {Ae : Equiv A} {Aap : Apart A} {Be : Equiv B} {Bap : Apart B} (f : A → B) `{inv : !Inverse f}. Class StrongInjective : Prop := { strong_injective : ∀ x y, x ≶ y → f x ≶ f y ; strong_injective_mor : StrongSetoid_Morphism f }. Class Injective : Prop := { injective : ∀ x y, f x = f y → x = y ; injective_mor : Setoid_Morphism f }. Class Surjective : Prop := { surjective : f ∘ (f ⁻¹) = id (* a.k.a. "split-epi" *) ; surjective_mor : Setoid_Morphism f }. Class Bijective : Prop := { bijective_injective :: Injective ; bijective_surjective :: Surjective }. End jections. #[global] Instance: Params (@StrongInjective) 4 := {}. #[global] Instance: Params (@Injective) 4 := {}. #[global] Instance: Params (@Surjective) 4 := {}. #[global] Instance: Params (@Bijective) 4 := {}. math-classes-8.19.0/interfaces/additional_operations.v000066400000000000000000000114301460576051100230660ustar00rootroot00000000000000Require Import Coq.Classes.Morphisms MathClasses.interfaces.abstract_algebra. Class Pow A B := pow : A → B → A. Infix "^" := pow : mc_scope. Notation "(^)" := pow (only parsing) : mc_scope. Notation "( x ^)" := (pow x) (only parsing) : mc_scope. Notation "(^ n )" := (λ x, x ^ n) (only parsing) : mc_scope. #[global] Instance: Params (@pow) 3 := {}. (* If we make [nat_pow_proper] a subclass, Coq is unable to find it. However, if we make a global instance in theory.nat_pow, it works? *) Class NatPowSpec A B (pw : Pow A B) `{Equiv A} `{Equiv B} `{One A} `{Mult A} `{Zero B} `{One B} `{Plus B} := { nat_pow_proper : Proper ((=) ==> (=) ==> (=)) (^) ; nat_pow_0 : ∀ x, x ^ 0 = 1 ; nat_pow_S : ∀ x n, x ^ (1 + n) = x * x ^ n }. Class IntPowSpec A B (pow : Pow A B) `{Equiv A} `{Equiv B} `{Zero A} `{One A} `{Mult A} `{Zero B} `{One B} `{Plus B} := { int_pow_proper : Proper ((=) ==> (=) ==> (=)) (^) ; int_pow_0 : ∀ x, x ^ 0 = 1 ; int_pow_base_0 : ∀ (n : B), n ≠ 0 → 0 ^ n = 0 ; int_pow_S : ∀ x n, x ≠ 0 → x ^ (1 + n) = x * x ^ n }. Class ShiftL A B := shiftl: A → B → A. Infix "≪" := shiftl (at level 33, left associativity) : mc_scope. Notation "(≪)" := shiftl (only parsing) : mc_scope. Notation "( x ≪)" := (shiftl x) (only parsing) : mc_scope. Notation "(≪ n )" := (λ x, x ≪ n) (only parsing) : mc_scope. #[global] Instance: Params (@shiftl) 3 := {}. Class ShiftLSpec A B (sl : ShiftL A B) `{Equiv A} `{Equiv B} `{One A} `{Plus A} `{Mult A} `{Zero B} `{One B} `{Plus B} := { shiftl_proper : Proper ((=) ==> (=) ==> (=)) (≪) ; shiftl_0 :: RightIdentity (≪) 0 ; shiftl_S : ∀ x n, x ≪ (1 + n) = 2 * x ≪ n }. Lemma shiftl_spec_from_nat_pow `{SemiRing A} `{SemiRing B} `{!NatPowSpec A B pw} (sl : ShiftL A B) : (∀ x n, x ≪ n = x * 2 ^ n) → ShiftLSpec A B sl. Proof. pose proof nat_pow_proper. intros spec. split. intros ? ? E1 ? ? E2. rewrite 2!spec. now rewrite E1, E2. intro x. rewrite spec, nat_pow_0. now apply right_identity. intros x n. rewrite 2!spec. rewrite nat_pow_S. now rewrite ?associativity, (commutativity x 2). Qed. Lemma shiftl_spec_from_int_pow `{SemiRing A} `{!PropHolds ((2:A) ≠ 0)} `{SemiRing B} `{!IntPowSpec A B ip} (sl : ShiftL A B) : (∀ x n, x ≪ n = x * 2 ^ n) → ShiftLSpec A B sl. Proof. pose proof int_pow_proper. intros spec. split. intros ? ? E1 ? ? E2. rewrite 2!spec. now rewrite E1, E2. intro x. rewrite spec, int_pow_0. now apply right_identity. intros x n. rewrite 2!spec. rewrite int_pow_S by solve_propholds. now rewrite ?associativity, (commutativity x 2). Qed. Class ShiftR A B := shiftr: A → B → A. Infix "≫" := shiftr (at level 33, left associativity) : mc_scope. Notation "(≫)" := shiftr (only parsing) : mc_scope. #[global] Instance: Params (@shiftr) 3 := {}. Class ShiftRSpec A B (sl : ShiftR A B) `{Equiv A} `{Equiv B} `{One A} `{Plus A} `{Mult A} `{Zero B} `{One B} `{Plus B} := { shiftr_proper : Proper ((=) ==> (=) ==> (=)) (≫) ; shiftr_0 :: RightIdentity (≫) 0 ; shiftr_S : ∀ x n, x ≫ n = 2 * x ≫ (1 + n) ∨ x ≫ n = 2 * x ≫ (1 + n) + 1 }. Class DivEuclid A := div_euclid : A → A → A. Class ModEuclid A := mod_euclid : A → A → A. Infix "`div`" := div_euclid (at level 35) : mc_scope. Notation "(`div`)" := div_euclid (only parsing) : mc_scope. Notation "( x `div`)" := (div_euclid x) (only parsing) : mc_scope. Notation "(`div` y )" := (λ x, x `div` y) (only parsing) : mc_scope. Infix "`mod`" := mod_euclid (at level 40) : mc_scope. Notation "(`mod` )" := mod_euclid (only parsing) : mc_scope. Notation "( x `mod`)" := (mod_euclid x) (only parsing) : mc_scope. Notation "(`mod` y )" := (λ x, x `mod` y) (only parsing) : mc_scope. #[global] Instance: Params (@div_euclid) 2 := {}. #[global] Instance: Params (@mod_euclid) 2 := {}. Class EuclidSpec A (d : DivEuclid A) (m : ModEuclid A) `{Equiv A} `{Le A} `{Lt A} `{Zero A} `{Plus A} `{Mult A} := { div_proper : Proper ((=) ==> (=) ==> (=)) (`div`) ; mod_proper : Proper ((=) ==> (=) ==> (=)) (`mod`) ; div_mod : ∀ x y, y ≠ 0 → x = y * x `div` y + x `mod` y ; mod_rem : ∀ x y, y ≠ 0 → 0 ≤ x `mod` y < y ∨ y < x `mod` y ≤ 0 ; div_0 : ∀ x, x `div` 0 = 0 ; mod_0 : ∀ x, x `mod` 0 = match (Nat.modulo 1 0) with | 0 => 0 | _ => x end }. Class CutMinus A := cut_minus : A → A → A. Infix "∸" := cut_minus (at level 50, left associativity) : mc_scope. Notation "(∸)" := cut_minus (only parsing) : mc_scope. Notation "( x ∸)" := (cut_minus x) (only parsing) : mc_scope. Notation "(∸ y )" := (λ x, x ∸ y) (only parsing) : mc_scope. #[global] Instance: Params (@cut_minus) 2 := {}. Class CutMinusSpec A (cm : CutMinus A) `{Equiv A} `{Zero A} `{Plus A} `{Le A} := { cut_minus_le : ∀ x y, y ≤ x → x ∸ y + y = x ; cut_minus_0 : ∀ x y, x ≤ y → x ∸ y = 0 }. math-classes-8.19.0/interfaces/canonical_names.v000066400000000000000000000501141460576051100216270ustar00rootroot00000000000000Global Generalizable All Variables. Global Set Automatic Introduction. Require Import MathClasses.theory.CoqStreams. Require Export Coq.Classes.Morphisms Coq.Setoids.Setoid Coq.Program.Program Coq.Unicode.Utf8 Coq.Unicode.Utf8_core MathClasses.misc.stdlib_hints. Definition id {A : Type} (a : A) := a. (* Equality *) Class Equiv A := equiv: relation A. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) (* Revert to transparency to allow conversions during unification. *) #[global] Typeclasses Transparent Equiv. #[global] Typeclasses Transparent compose flip. Set Warnings "+unsupported-attributes". (* We use this virtually everywhere, and so use "=" for it: *) Infix "=" := equiv : type_scope. Notation "(=)" := equiv (only parsing) : mc_scope. Notation "( x =)" := (equiv x) (only parsing) : mc_scope. Notation "(= x )" := (λ y, equiv y x) (only parsing) : mc_scope. Notation "(≠)" := (λ x y, ¬x = y) (only parsing) : mc_scope. Notation "x ≠ y":= (¬x = y): type_scope. Notation "( x ≠)" := (λ y, x ≠ y) (only parsing) : mc_scope. Notation "(≠ x )" := (λ y, y ≠ x) (only parsing) : mc_scope. Delimit Scope mc_scope with mc. Global Open Scope mc_scope. #[global] Hint Extern 2 (?x = ?x) => reflexivity. #[global] Hint Extern 2 (?x = ?y) => auto_symm. #[global] Hint Extern 2 (?x = ?y) => auto_trans. (* Coq sometimes uses an incorrect DefaultRelation, so we override it. *) #[global] Instance equiv_default_relation `{Equiv A} : DefaultRelation (=) | 3 := {}. (* Because Coq does not support setoid rewriting in Type'd relations we put apartness in Prop. This has as obvious disadvantage, that when an implementation actually needs the witness from the apartness relation (e.g. for division), it has to use something like [constructive_indefinite_description] to extract it. This is quite inefficient and will also most likely block due to proofs being Opaque. Nonetheless we do not think this is a problem because for actual computation, one does not want to rely on some arbitrary witness obtained by chaining a lot of proofs. Instead one would (and probably should) specify an optimal witness by hand. Another advantage of our approach is that classes describing structures (e.g. Field) can remain in Prop. *) Class Apart A := apart: relation A. Infix "≶" := apart (at level 70, no associativity) : mc_scope. Notation "(≶)" := apart (only parsing) : mc_scope. Notation "( x ≶)" := (apart x) (only parsing) : mc_scope. Notation "(≶ x )" := (λ y, apart y x) (only parsing) : mc_scope. (* Even for setoids with decidable equality x ≠ y does not imply x ≶ y. Therefore we introduce the following class. *) Class TrivialApart A `{Equiv A} {Aap : Apart A} := trivial_apart : ∀ x y, x ≶ y ↔ x ≠ y. (* For Leibniz equality we use "≡", We do not define it as setoid equality with low priority because sometimes we are interested in both setoid and Leibniz equality on the same structure. *) Infix "≡" := eq (at level 70, no associativity) : mc_scope. Notation "(≡)" := eq (only parsing) : mc_scope. Notation "( x ≡)" := (eq x) (only parsing) : mc_scope. Notation "(≡ x )" := (λ y, eq y x) (only parsing) : mc_scope. Notation "(≢)" := (λ x y, ¬x ≡ y) (only parsing) : mc_scope. Notation "x ≢ y":= (¬x ≡ y) (at level 70, no associativity) : mc_scope. Notation "( x ≢)" := (λ y, x ≢ y) (only parsing) : mc_scope. Notation "(≢ x )" := (λ y, y ≢ x) (only parsing) : mc_scope. (* Some common notions of equality *) Definition ext_equiv `{Equiv A} `{Equiv B} : Equiv (A → B) := ((=) ==> (=))%signature. #[global] Hint Extern 10 (Equiv (_ → _)) => apply @ext_equiv : typeclass_instances. #[global] Hint Extern 10 (Equiv (relation _)) => apply @ext_equiv : typeclass_instances. (* Due to bug #2491 *) (** Interestingly, most of the development works fine if this is defined as ∀ x, f x = g x. However, in the end that version was just not strong enough for comfortable rewriting in setoid-pervasive contexts. *) Notation "x ↾ p" := (exist _ x p) (at level 20) : mc_scope. Definition sig_equiv `{Equiv A} (P: A → Prop) : Equiv (sig P) := λ x y, `x = `y. Ltac simpl_sig_equiv := match goal with | |- (@equiv _ (@sig_equiv _ ?e _) (?x↾_) (?y↾_)) => change (@equiv _ e x y) end. #[global] Hint Extern 10 (Equiv (sig _)) => apply @sig_equiv : typeclass_instances. #[global] Hint Extern 4 (@equiv _ (sig_equiv _ _ _) (_↾_) (_↾_)) => simpl_sig_equiv. Definition sigT_equiv `{Equiv A} (P: A → Type) : Equiv (sigT P) := λ a b, projT1 a = projT1 b. #[global] Hint Extern 10 (Equiv (sigT _)) => apply @sigT_equiv : typeclass_instances. Definition sig_apart `{Apart A} (P: A → Prop) : Apart (sig P) := λ x y, `x ≶ `y. #[global] Hint Extern 10 (Apart (sig _)) => apply @sig_apart : typeclass_instances. Class Cast A B := cast: A → B. Arguments cast _ _ {Cast} _. Notation "' x" := (cast _ _ x) (at level 20) : mc_scope. #[global] Instance: Params (@cast) 3 := {}. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Transparent Cast. Set Warnings "+unsupported-attributes". (* Other canonically named relations/operations/constants: *) Class SgOp A := sg_op: A → A → A. Class MonUnit A := mon_unit: A. Class Plus A := plus: A → A → A. Class Mult A := mult: A → A → A. Class One A := one: A. Class Zero A := zero: A. Class Negate A := negate: A → A. Class DecRecip A := dec_recip: A → A. Definition ApartZero R `{Zero R} `{Apart R} := sig (≶ zero). Class Recip A `{Apart A} `{Zero A} := recip: ApartZero A → A. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Transparent SgOp MonUnit Plus Mult Zero One Negate. Set Warnings "+unsupported-attributes". Class Meet A := meet: A → A → A. Class Join A := join: A → A → A. Class Top A := top: A. Class Bottom A := bottom: A. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Transparent Meet Join Top Bottom. Set Warnings "+unsupported-attributes". Class Contains A B := contains: A → B → Prop. Class Singleton A B := singleton: A → B. Class Difference A := difference : A → A → A. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Transparent Contains Singleton Difference. Set Warnings "+unsupported-attributes". Class Le A := le: relation A. Class Lt A := lt: relation A. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Transparent Le Lt. Set Warnings "+unsupported-attributes". Definition NonNeg R `{Zero R} `{Le R} := sig (le zero). Definition Pos R `{Zero R} `{Equiv R} `{Lt R} := sig (lt zero). Definition NonPos R `{Zero R} `{Le R} := sig (λ y, le y zero). Inductive PosInf (R : Type) : Type := finite (x : R) | infinite. Class Arrows (O: Type): Type := Arrow: O → O → Type. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Transparent Arrows. (* Ideally this should be removed *) Set Warnings "+unsupported-attributes". Infix "⟶" := Arrow (at level 90, right associativity) : mc_scope. Class CatId O `{Arrows O} := cat_id: ∀ x, x ⟶ x. Class CatComp O `{Arrows O} := comp: ∀ x y z, (y ⟶ z) → (x ⟶ y) → (x ⟶ z). Class RalgebraAction A B := ralgebra_action: A → B → B. Arguments cat_id {O arrows CatId x} : rename. Arguments comp {O arrow CatComp} _ _ _ _ _ : rename. #[global] Instance: Params (@mult) 2 := {}. #[global] Instance: Params (@plus) 2 := {}. #[global] Instance: Params (@negate) 2 := {}. #[global] Instance: Params (@equiv) 2 := {}. #[global] Instance: Params (@apart) 2 := {}. #[global] Instance: Params (@le) 2 := {}. #[global] Instance: Params (@lt) 2 := {}. #[global] Instance: Params (@recip) 4 := {}. #[global] Instance: Params (@dec_recip) 2 := {}. #[global] Instance: Params (@meet) 2 := {}. #[global] Instance: Params (@join) 2 := {}. #[global] Instance: Params (@contains) 3 := {}. #[global] Instance: Params (@singleton) 3 := {}. #[global] Instance: Params (@difference) 2 := {}. #[global] Instance plus_is_sg_op `{f : Plus A} : SgOp A := f. #[global] Instance mult_is_sg_op `{f : Mult A} : SgOp A := f. #[global] Instance one_is_mon_unit `{c : One A} : MonUnit A := c. #[global] Instance zero_is_mon_unit `{c : Zero A} : MonUnit A := c. #[global] Instance meet_is_sg_op `{f : Meet A} : SgOp A := f. #[global] Instance join_is_sg_op `{f : Join A} : SgOp A := f. #[global] Instance top_is_mon_unit `{s : Top A} : MonUnit A := s. #[global] Instance bottom_is_mon_unit `{s : Bottom A} : MonUnit A := s. #[global] Instance singleton_is_cast `{s : Singleton A B} : Cast A B := s. #[global] Hint Extern 10 (Equiv (_ ⟶ _)) => apply @ext_equiv : typeclass_instances. #[global] Hint Extern 4 (Equiv (ApartZero _)) => apply @sig_equiv : typeclass_instances. #[global] Hint Extern 4 (Equiv (NonNeg _)) => apply @sig_equiv : typeclass_instances. #[global] Hint Extern 4 (Equiv (Pos _)) => apply @sig_equiv : typeclass_instances. #[global] Hint Extern 4 (Equiv (PosInf _)) => apply @sig_equiv : typeclass_instances. #[global] Hint Extern 4 (Apart (ApartZero _)) => apply @sig_apart : typeclass_instances. #[global] Hint Extern 4 (Apart (NonNeg _)) => apply @sig_apart : typeclass_instances. #[global] Hint Extern 4 (Apart (Pos _)) => apply @sig_apart : typeclass_instances. #[global] Hint Extern 4 (Apart (PosInf _)) => apply @sig_apart : typeclass_instances. (* Notations: *) Notation "R ₀" := (ApartZero R) (at level 20, no associativity) : mc_scope. Notation "R ⁺" := (NonNeg R) (at level 20, no associativity) : mc_scope. Notation "R ₊" := (Pos R) (at level 20, no associativity) : mc_scope. Notation "R ⁻" := (NonPos R) (at level 20, no associativity) : mc_scope. Notation "R ∞" := (PosInf R) (at level 20, no associativity) : mc_scope. Infix "&" := sg_op (at level 50, left associativity) : mc_scope. Notation "(&)" := sg_op (only parsing) : mc_scope. Notation "( x &)" := (sg_op x) (only parsing) : mc_scope. Notation "(& x )" := (λ y, y & x) (only parsing) : mc_scope. Infix "+" := plus : mc_scope. Notation "(+)" := plus (only parsing) : mc_scope. Notation "( x +)" := (plus x) (only parsing) : mc_scope. Notation "(+ x )" := (λ y, y + x) (only parsing) : mc_scope. Infix "*" := mult : mc_scope. (* We don't add "( * )", "( * x )" and "( x * )" notations because they conflict with comments. *) Notation "( x *.)" := (mult x) (only parsing) : mc_scope. Notation "(.*.)" := mult (only parsing) : mc_scope. Notation "(.* x )" := (λ y, y * x) (only parsing) : mc_scope. Notation "- x" := (negate x) : mc_scope. Notation "(-)" := negate (only parsing) : mc_scope. Notation "x - y" := (x + -y) : mc_scope. Notation "0" := zero : mc_scope. Notation "1" := one : mc_scope. Notation "2" := (1 + 1) : mc_scope. Notation "3" := (1 + (1 + 1)) : mc_scope. Notation "4" := (1 + (1 + (1 + 1))) : mc_scope. Notation "- 1" := (-(1)) : mc_scope. Notation "- 2" := (-(2)) : mc_scope. Notation "- 3" := (-(3)) : mc_scope. Notation "- 4" := (-(4)) : mc_scope. Notation "/ x" := (dec_recip x) : mc_scope. Notation "(/)" := dec_recip (only parsing) : mc_scope. Notation "x / y" := (x * /y) : mc_scope. Notation "// x" := (recip x) (at level 35, right associativity) : mc_scope. Notation "(//)" := recip (only parsing) : mc_scope. Notation "x // y" := (x * //y) (at level 35, right associativity) : mc_scope. Notation "⊤" := top : mc_scope. Notation "⊥" := bottom : mc_scope. Infix "⊓" := meet (at level 50, no associativity) : mc_scope. Notation "(⊓)" := meet (only parsing) : mc_scope. Notation "( X ⊓)" := (meet X) (only parsing) : mc_scope. Notation "(⊓ X )" := (λ Y, Y ⊓ X) (only parsing) : mc_scope. Infix "⊔" := join (at level 50, no associativity) : mc_scope. Notation "(⊔)" := join (only parsing) : mc_scope. Notation "( X ⊔)" := (join X) (only parsing) : mc_scope. Notation "(⊔ X )" := (λ Y, Y ⊔ X) (only parsing) : mc_scope. Infix "≤" := le : mc_scope. Notation "(≤)" := le (only parsing) : mc_scope. Notation "( x ≤)" := (le x) (only parsing) : mc_scope. Notation "(≤ x )" := (λ y, y ≤ x) (only parsing) : mc_scope. Infix "<" := lt : mc_scope. Notation "(<)" := lt (only parsing) : mc_scope. Notation "( x <)" := (lt x) (only parsing) : mc_scope. Notation "(< x )" := (λ y, y < x) (only parsing) : mc_scope. Notation "x ≤ y ≤ z" := (x ≤ y ∧ y ≤ z) (at level 70, y at next level) : mc_scope. Notation "x ≤ y < z" := (x ≤ y ∧ y < z) (at level 70, y at next level) : mc_scope. Notation "x < y < z" := (x < y ∧ y < z) (at level 70, y at next level) : mc_scope. Notation "x < y ≤ z" := (x < y ∧ y ≤ z) (at level 70, y at next level) : mc_scope. Infix "∖" := difference (at level 35) : mc_scope. Notation "(∖)" := difference (only parsing) : mc_scope. Notation "( X ∖)" := (difference X) (only parsing) : mc_scope. Notation "(∖ X )" := (λ Y, Y ∖ X) (only parsing) : mc_scope. Infix "∈" := contains (at level 70, no associativity) : mc_scope. Notation "(∈)" := contains (only parsing) : mc_scope. Notation "( x ∈)" := (contains x) (only parsing) : mc_scope. Notation "(∈ X )" := (λ x, x ∈ X) (only parsing) : mc_scope. Notation "x ∉ y" := (¬x ∈ y) (at level 70, no associativity) : mc_scope. Notation "(∉)" := (λ x X, x ∉ X) : mc_scope. Notation "( x ∉)" := (λ X, x ∉ X) (only parsing) : mc_scope. Notation "(∉ X )" := (λ x, x ∉ X) (only parsing) : mc_scope. Notation "{{ x }}" := (singleton x) : mc_scope. Notation "{{ x ; y ; .. ; z }}" := (join .. (join (singleton x) (singleton y)) .. (singleton z)) : mc_scope. Infix "◎" := (comp _ _ _) (at level 40, left associativity) : mc_scope. (* Taking over ∘ is just a little too zealous at this point. With our current approach, it would require changing all (nondependent) function types A → B with A ⟶ B to make them use the canonical name for arrows, which is a tad extreme. *) Notation "(◎)" := (comp _ _ _) (only parsing) : mc_scope. Notation "( f ◎)" := (comp _ _ _ f) (only parsing) : mc_scope. Notation "(◎ f )" := (λ g, comp _ _ _ g f) (only parsing) : mc_scope. (* Haskell style! *) Notation "(→)" := (λ x y, x → y) : mc_scope. Notation "t $ r" := (t r) (at level 65, right associativity, only parsing) : mc_scope. Notation "(∘)" := compose (only parsing) : mc_scope. (* Agda style! *) Notation "∞ X" := (Stream X) (at level 23) : mc_scope. Infix ":::" := Cons (at level 60, right associativity) : mc_scope. Notation "(:::)" := Cons (only parsing) : mc_scope. Notation "(::: X )" := (λ x, Cons x X) (only parsing) : mc_scope. Notation "( x :::)" := (Cons x) (only parsing) : mc_scope. #[global] Hint Extern 2 (?x ≤ ?y) => reflexivity. #[global] Hint Extern 4 (?x ≤ ?z) => auto_trans. #[global] Hint Extern 4 (?x < ?z) => auto_trans. Class Abs A `{Equiv A} `{Le A} `{Zero A} `{Negate A} := abs_sig: ∀ (x : A), { y : A | (0 ≤ x → y = x) ∧ (x ≤ 0 → y = -x)}. Definition abs `{Abs A} := λ x : A, ` (abs_sig x). #[global] Instance: Params (@abs_sig) 6 := {}. #[global] Instance: Params (@abs) 6 := {}. (* Common properties: *) Class Inverse `(A → B) : Type := inverse: B → A. Arguments inverse {A B} _ {Inverse} _. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Transparent Inverse. Set Warnings "+unsupported-attributes". Notation "f ⁻¹" := (inverse f) (at level 30) : mc_scope. Class Idempotent `{ea : Equiv A} (f: A → A → A) (x : A) : Prop := idempotency: f x x = x. Arguments idempotency {A ea} _ _ {Idempotent}. Class UnaryIdempotent `{Equiv A} (f: A → A) : Prop := unary_idempotent :: Idempotent (@compose A A A) f. Lemma unary_idempotency `{Equiv A} `{!Reflexive (=)} `{!UnaryIdempotent f} x : f (f x) = f x. Proof. firstorder. Qed. Class BinaryIdempotent `{Equiv A} (op: A → A → A) : Prop := binary_idempotent :: ∀ x, Idempotent op x. Class LeftIdentity {A} `{Equiv B} (op : A → B → B) (x : A): Prop := left_identity: ∀ y, op x y = y. Class RightIdentity `{Equiv A} {B} (op : A → B → A) (y : B): Prop := right_identity: ∀ x, op x y = x. Class Absorption `{Equiv A} {B} {C} (op1: A → C → A) (op2: A → B → C) : Prop := absorption: ∀ x y, op1 x (op2 x y) = x. Class LeftAbsorb `{Equiv A} {B} (op : A → B → A) (x : A): Prop := left_absorb: ∀ y, op x y = x. Class RightAbsorb {A} `{Equiv B} (op : A → B → B) (y : B): Prop := right_absorb: ∀ x, op x y = y. Class LeftInverse {A} {B} `{Equiv C} (op : A → B → C) (inv : B → A) (unit : C) := left_inverse: ∀ x, op (inv x) x = unit. Class RightInverse {A} {B} `{Equiv C} (op : A → B → C) (inv : A → B) (unit : C) := right_inverse: ∀ x, op x (inv x) = unit. Class Commutative `{Equiv B} `(f : A → A → B) : Prop := commutativity: ∀ x y, f x y = f y x. Class HeteroAssociative {A B C AB BC} `{Equiv ABC} (fA_BC: A → BC → ABC) (fBC: B → C → BC) (fAB_C: AB → C → ABC) (fAB : A → B → AB): Prop := associativity : ∀ x y z, fA_BC x (fBC y z) = fAB_C (fAB x y) z. Class Associative `{Equiv A} f := simple_associativity:: HeteroAssociative f f f f. Notation ArrowsAssociative C := (∀ {w x y z: C}, HeteroAssociative (◎) (comp z _ _ ) (◎) (comp y x w)). Class Involutive `{Equiv A} (f : A → A) := involutive: ∀ x, f (f x) = x. Class TotalRelation `(R : relation A) : Prop := total : ∀ x y : A, R x y ∨ R y x. Arguments total {A} _ {TotalRelation} _ _. Class Trichotomy `{Ae : Equiv A} `(R : relation A) := trichotomy : ∀ x y : A, R x y ∨ x = y ∨ R y x. Arguments trichotomy {A Ae} _ {Trichotomy} _ _. Arguments irreflexivity {A} _ {Irreflexive} _ _. Class CoTransitive `(R : relation A) : Prop := cotransitive : ∀ x y, R x y → ∀ z, R x z ∨ R z y. Arguments cotransitive {A R CoTransitive x y} _ _. Class AntiSymmetric `{Ae : Equiv A} (R : relation A) : Prop := antisymmetry: ∀ x y, R x y → R y x → x = y. Arguments antisymmetry {A Ae} _ {AntiSymmetric} _ _ _ _. Class LeftHeteroDistribute {A B} `{Equiv C} (f : A → B → C) (g_r : B → B → B) (g : C → C → C) : Prop := distribute_l : ∀ a b c, f a (g_r b c) = g (f a b) (f a c). Class RightHeteroDistribute {A B} `{Equiv C} (f : A → B → C) (g_l : A → A → A) (g : C → C → C) : Prop := distribute_r: ∀ a b c, f (g_l a b) c = g (f a c) (f b c). Class LeftDistribute`{Equiv A} (f g: A → A → A) := simple_distribute_l :: LeftHeteroDistribute f g g. Class RightDistribute `{Equiv A} (f g: A → A → A) := simple_distribute_r :: RightHeteroDistribute f g g. Class HeteroSymmetric {A} {T : A → A → Type} (R : ∀ {x y}, T x y → T y x → Prop) : Prop := hetero_symmetric `(a : T x y) (b : T y x) : R a b → R b a. (* Although cancellation is the same as being injective, we want a proper name to refer to this commonly used property. *) Section cancellation. Context `{Ae : Equiv A} `{Aap : Apart A} (op : A → A → A) (z : A). Class LeftCancellation := left_cancellation : ∀ x y, op z x = op z y → x = y. Class RightCancellation := right_cancellation : ∀ x y, op x z = op y z → x = y. Class StrongLeftCancellation := strong_left_cancellation : ∀ x y, x ≶ y → op z x ≶ op z y. Class StrongRightCancellation := strong_right_cancellation : ∀ x y, x ≶ y → op x z ≶ op y z. End cancellation. (* Common names for properties that hold in N, Z, Q, ... *) Class ZeroProduct A `{Equiv A} `{!Mult A} `{!Zero A} : Prop := zero_product : ∀ x y, x * y = 0 → x = 0 ∨ y = 0. Class ZeroDivisor {R} `{Equiv R} `{Zero R} `{Mult R} (x : R) : Prop := zero_divisor : x ≠ 0 ∧ ∃ y, y ≠ 0 ∧ x * y = 0. Class NoZeroDivisors R `{Equiv R} `{Zero R} `{Mult R} : Prop := no_zero_divisors x : ¬ZeroDivisor x. #[global] Instance zero_product_no_zero_divisors `{ZeroProduct A} : NoZeroDivisors A. Proof. intros x [? [? [? E]]]. destruct (zero_product _ _ E); intuition. Qed. Class RingUnit `{Equiv R} `{Mult R} `{One R} (x : R) : Prop := ring_unit : ∃ y, x * y = 1. (* A common induction principle for both the naturals and integers *) Class Biinduction R `{Equiv R} `{Zero R} `{One R} `{Plus R} : Prop := biinduction (P : R → Prop) `{!Proper ((=) ==> iff) P} : P 0 → (∀ n, P n ↔ P (1 + n)) → ∀ n, P n. math-classes-8.19.0/interfaces/finite_sets.v000066400000000000000000000071551460576051100210400ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders. (* We define finite sets as the initial join semi lattice over a decidable type. An important aspect of our interface is the possibility to make fast implementations. Therefore it is important that various properties of the carrier set can be used to create an implementation. Examples of implementations are: * Finite sets as unordered lists. * Finite sets as AVLs (here we need order). * Finite sets of positives as trees with boolean nodes. The above examples already illustrate that it is very inconvenient to fix the properties of the carrier right now. Instead we use the following type class: *) Class SetType (A : Type) := set_type: Type. Arguments set_type _ {SetType}. (* We can't make this type transparent to typeclass resolution. *) (* The key idea is that an implementation can let instances of this class depend on any specific properties it likes. For convenience we define some notations to hide nasty details. *) Notation EmptySet A := (Bottom (set_type A)). Notation "∅" := (@bottom (set_type _) _) : mc_scope. Notation SetEquiv A := (Equiv (set_type A)). Notation SetJoin A := (Join (set_type A)). Notation SetMeet A := (Meet (set_type A)). Notation SetDifference A := (Difference (set_type A)). Notation SetSingleton A := (Singleton A (set_type A)). Notation SetLe A := (Le (set_type A)). Notation SetContains A := (Contains A (set_type A)). Class FSetExtend A `{t : SetType A} := fset_extend `{Join B} `{Bottom B} : (A → B) → set_type A → B. Class FSet A `{At : SetType A} `{Ae : Equiv A} `{Ate : SetEquiv A} `{Aempty : EmptySet A} `{Ajoin : SetJoin A} `{Asingle : SetSingleton A} `{∀ a₁ a₂ : A, Decision (a₁ = a₂)} `{U : !FSetExtend A} := { fset_bounded_sl :: BoundedJoinSemiLattice (set_type A) ; singleton_mor :: Setoid_Morphism singleton ; fset_extend_mor `{BoundedJoinSemiLattice B} `{!Setoid_Morphism (f : A → B)} :: BoundedJoinSemiLattice_Morphism (fset_extend f) ; fset_extend_correct `{BoundedJoinSemiLattice B} (f : A → B) `{!Setoid_Morphism f} : f = fset_extend f ∘ singleton ; fset_extend_unique `{Equiv B} `{Join B} `{Bottom B} (f : A → B) `{!Setoid_Morphism f} (h : set_type A → B) `{!BoundedJoinSemiLattice_Morphism h} : f = h ∘ singleton → h = fset_extend f }. Definition fset_map `(f : A → B) `{SetType A} `{SetType B} `{EmptySet B} `{SetJoin B} `{SetSingleton B} `{U : !FSetExtend A} : set_type A → set_type B := fset_extend (singleton ∘ f). (* Next we describe the order and set inclusion. Constructing a JoinSemiLatticeOrder might be quite inconvenient since we have used the algebraic definition of a lattice in FSet. However, an implementation can freely use orders.lattices.alt_Build_JoinSemiLatticeOrder. *) Class FSetContainsSpec A `{At : SetType A} `{Ae : Equiv A} `{Ate : SetEquiv A} `{SetLe A} `{SetContains A} `{Ajoin : SetJoin A} `{Asingle : SetSingleton A} := { fset_join_sl_order :: JoinSemiLatticeOrder (≤) ; fset_in_singleton_le : ∀ x X, x ∈ X ↔ {{ x }} ≤ X }. (* Unfortunately, properties as meet and the differences cannot be uniquely defined in an algebraic way, therefore we just use set inclusion. *) Class FullFSet A {car Ae conAe conAle Acontains Aempty Ajoin Asingle U Adec} `{Adiff : SetDifference A} `{Ameet : SetMeet A} := { full_fset_fset :: @FSet A car Ae conAe Aempty Ajoin Asingle U Adec ; full_fset_contains :: @FSetContainsSpec A car Ae conAe conAle Acontains Ajoin Asingle ; fset_in_meet : ∀ X Y x, x ∈ X ⊓ Y ↔ (x ∈ X ∧ x ∈ Y) ; fset_in_difference : ∀ X Y x, x ∈ X∖ Y ↔ (x ∈ X ∧ x ∉ Y) }. math-classes-8.19.0/interfaces/functors.v000066400000000000000000000130661460576051100203650ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra. Require MathClasses.theory.setoids. Section functor_class. Context `{Category C} `{Category D} (M: C → D). Class Fmap: Type := fmap: ∀ {v w: C}, (v ⟶ w) → (M v ⟶ M w). Class Functor `(Fmap): Prop := { functor_from: Category C ; functor_to: Category D ; functor_morphism:: ∀ a b: C, Setoid_Morphism (@fmap _ a b) ; preserves_id: `(fmap (cat_id: a ⟶ a) = cat_id) ; preserves_comp `(f: y ⟶ z) `(g: x ⟶ y): fmap (f ◎ g) = fmap f ◎ fmap g }. End functor_class. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Transparent Fmap. Set Warnings "+unsupported-attributes". (* The usual notational convention for functor application is to use the name of the functor to refer to both its object map and its arrow map, relying on additional conventions regarding object/arrow names for disambiguation (i.e. "F x" and "F f" map an object and an arrow, respectively, because "x" and "f" are conventional names for objects and arrows, respectively. In Coq, for a term F to function as though it had two different types simultaneously (namely the object map and the arrow map), there must either (1) be coercions from the type of F to either function, or (2) F must be (coercible to) a single function that is able to consume both object and arrow arguments. The following snippet shows that (1) doesn't appear to be supported in Coq: Section test. Context (A B: Type). Record R := { Ra:> A → unit; Rb:> B → unit }. Variables (r: R) (a: A) (b: B). Check (r a). (* ok so far *) Check (r b). (* Error: The term "b" has type "B" while it is expected to have type "A". *) End test. And even if this /did/ work, we could not use it, because leaving computational components unbundled is a key aspect of our approach. For (2), if it could be made to work at all (which is not clear at all), F would need a pretty egregious type considering that arrow types are indexed by objects, and that the type of the arrow map (namely "∀ x y, (x ⟶ y) → (F x ⟶ F y)") must refer to the object map. We feel that these issues are not limitations of the Coq system, but merely reflect the fact that notationally identifying these two different and interdependent maps is a typical example of an "abus de notation" that by its very nature is ill-suited to a formal development where software engineering concerns apply. Hence, we do not adopt this practice, and use "fmap F" (name taken from the Haskell standard library) to refer to the arrow map of a functor F. TODO: Sharpen. *) Section id_functor. Context `{Category C}. Global Instance: Fmap id := λ _ _, id. Global Instance id_functor: Functor (id: C → C) _. Proof. constructor; try reflexivity; try apply _. Qed. End id_functor. Section compose_functors. Context A B C `{!Arrows A} `{!Arrows B} `{!Arrows C} `{!CatId A} `{!CatId B} `{!CatId C} `{!CatComp A} `{!CatComp B} `{!CatComp C} `{∀ x y: A, Equiv (x ⟶ y)} `{∀ x y: B, Equiv (x ⟶ y)} `{∀ x y: C, Equiv (x ⟶ y)} `{!Functor (f: B → C) f'} `{!Functor (g: A → B) g'}. Global Instance comp_Fmap: Fmap (f ∘ g) := λ _ _, fmap f ∘ fmap g. Global Instance compose_functors: Functor (f ∘ g) _. Proof with intuition; try apply _. pose proof (functor_from g). pose proof (functor_to g). pose proof (functor_to f). constructor; intros; try apply _. apply (@setoids.compose_setoid_morphism _ _ _ _ _ _)... apply (@functor_morphism _ _ _ _ _ _ _ _ _ _ f _)... (* todo: this part really should be automatic *) change (fmap f (fmap g (cat_id: a ⟶ a)) = cat_id). repeat try rewrite preserves_id... change (fmap f (fmap g (f0 ◎ g0)) = fmap f (fmap g f0) ◎ fmap f (fmap g g0)). repeat try rewrite preserves_comp... Qed. End compose_functors. (** The Functor class is nice and abstract and theory-friendly, but not as convenient to use for regular programming as Haskell's Functor class. The reason for this is that our Functor is parameterized on two Categories, which by necessity bundle setoid- ness and setoid-morphism-ness into objects and arrows, respectively. The Haskell Functor class, by contrast, is essentially specialized for endofunctors on the category of Haskell types and functions between them. The latter corresponds to our setoid.Object category. To recover convenience, we introduce a second functor type class tailored specifically to functors of this kind. The specialization allows us to forgo bundling, and lets us recover the down-to-earth Type→Type type for the type constructor, and the (a→b)→(F a→F b) type for the function map, with all setoid/morphism proofs hidden in the structure class in Prop. To justify this definition, in theory/functors we show that instances of this new functor class do indeed give rise to instances of the original nice abstract Functor class. *) Class SFmap (M : Type → Type) := sfmap: ∀ `(A → B), (M A → M B). Class SFunctor (M : Type → Type) `{∀ `{Equiv A}, Equiv (M A)} `{SFmap M} : Prop := { sfunctor_setoid `{Setoid A} :: Setoid (M A) ; sfmap_proper `{Setoid A} `{Setoid B} :: Proper (((=) ==> (=)) ==> ((=) ==> (=))) (@sfmap M _ A B) ; sfmap_id `{Setoid A} : sfmap id = id ; sfmap_comp `{Equiv A} `{Equiv B} `{Equiv C} `{!Setoid_Morphism (f : B → C)} `{!Setoid_Morphism (g : A → B)} : sfmap (f ∘ g) = sfmap f ∘ sfmap g }. math-classes-8.19.0/interfaces/integers.v000066400000000000000000000041041460576051100203330ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals MathClasses.theory.categories MathClasses.categories.varieties. Require MathClasses.varieties.rings. Section initial_maps. Variable A : Type. Class IntegersToRing := integers_to_ring: ∀ R `{Mult R} `{Plus R} `{One R} `{Zero R} `{Negate R}, A → R. Context `{IntegersToRing} `{Ring A} `{∀ `{Ring B}, SemiRing_Morphism (integers_to_ring B)}. Global Instance integer_initial_arrow: InitialArrow (rings.object A). intro y. exists (λ u, match u return A → y u with tt => integers_to_ring (y tt) end). abstract (apply rings.encode_morphism_only; apply _). Defined. (* for some reason [Program] isn't cooperating here. look into it *) Lemma integer_initial (same_morphism : ∀ `{Ring B} {h : A → B} `{!SemiRing_Morphism h}, integers_to_ring B = h) : Initial (rings.object A). Proof. intros y [x h] [] ?. simpl in *. apply same_morphism. apply rings.decode_variety_and_ops. apply (@rings.decode_morphism_and_ops _ _ _ _ _ _ _ _ _ h). reflexivity. Qed. End initial_maps. #[global] Instance: Params (@integers_to_ring) 8 := {}. Class Integers A {e plus mult zero one negate} `{U : IntegersToRing A} := { integers_ring:: @Ring A e plus mult zero one negate ; integers_to_ring_mor:: ∀ `{Ring B}, SemiRing_Morphism (integers_to_ring A B) ; integers_initial:: Initial (rings.object A) }. Section specializable. Context (Z N : Type) `{Integers Z} `{Naturals N}. Class IntAbs := int_abs_sig : ∀ x, { n : N | naturals_to_semiring N Z n = x } + { n : N | naturals_to_semiring N Z n = -x }. Definition int_abs `{ia : IntAbs} (x : Z) : N := match int_abs_sig x with | inl (n↾_) => n | inr (n↾_) => n end. Definition int_to_nat `{Zero N} `{ia : IntAbs} (x : Z) : N := match int_abs_sig x with | inl (n↾_) => n | inr (n↾_) => 0 end. End specializable. #[global] Instance: Params (@int_abs) 10 := {}. #[global] Instance: Params (@int_abs_sig) 10 := {}. #[global] Instance: Params (@int_to_nat) 11 := {}. math-classes-8.19.0/interfaces/monads.v000066400000000000000000000047041460576051100200020ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.canonical_names. Require Export MathClasses.interfaces.functors. Section ops. Context (M : Type → Type). Class MonadReturn := ret: ∀ {A}, A → M A. Class MonadBind := bind: ∀ {A B}, (A → M B) → M A → M B. Class MonadJoin := join: ∀ {A}, M (M A) → M A. End ops. Arguments ret {M MonadReturn A} _. Arguments bind {M MonadBind A B} _ _. Arguments join {M MonadJoin A} _. Notation "m ≫= f" := (bind f m) (at level 60, right associativity) : mc_scope. Notation "x ← y ; z" := (y ≫= (λ x : _, z)) (at level 65, (*next at level 35, *) right associativity, only parsing) : mc_scope. (* Notation "x ≫ y" := (_ ← x ; y) (at level 33, right associativity, only parsing) : mc_scope. *) Class Monad (M : Type → Type) `{∀ A, Equiv A → Equiv (M A)} `{MonadReturn M} `{MonadBind M} : Prop := { mon_ret_proper `{Setoid A} :: Proper ((=) ==> (=)) (@ret _ _ A) ; mon_bind_proper `{Setoid A} `{Setoid B} :: Proper (((=) ==> (=)) ==> (=) ==> (=)) (@bind _ _ A B) ; mon_setoid `{Setoid A} :: Setoid (M A) ; bind_lunit `{Equiv A} `{Setoid B} `{!Setoid_Morphism (f : A → M B)} : bind f ∘ ret = f ; bind_runit `{Setoid A} : bind ret = id ; bind_assoc `{Equiv A} `{Equiv B} `{Setoid C} `{!Setoid_Morphism (f : B → M C)} `{!Setoid_Morphism (g : A → M B)} : bind f ∘ bind g = bind (bind f ∘ g) }. Class StrongMonad (M : Type → Type) `{∀ A, Equiv A → Equiv (M A)} `{MonadReturn M} `{SFmap M} `{MonadJoin M} : Prop := { smon_ret_proper `{Setoid A} :: Proper ((=) ==> (=)) (@ret _ _ A) ; smon_join_proper `{Setoid A} :: Proper ((=) ==> (=)) (@join _ _ A) ; smon_sfunctor :: SFunctor M ; sfmap_ret `{Equiv A} `{Equiv B} `{!Setoid_Morphism (f : A → B)} : sfmap f ∘ ret = ret ∘ f ; sfmap_join `{Equiv A} `{Equiv B} `{!Setoid_Morphism (f : A → B)} : sfmap f ∘ join = join ∘ sfmap (sfmap f) ; join_ret `{Setoid A} : join ∘ ret = id ; join_sfmap_ret `{Setoid A} : join ∘ sfmap ret = id ; join_sfmap_join `{Setoid A} : join ∘ sfmap join = join ∘ join }. Class FullMonad (M : Type → Type) `{∀ A, Equiv A → Equiv (M A)} `{MonadReturn M} `{MonadBind M} `{SFmap M} `{MonadJoin M} : Prop := { full_mon_mon :: Monad M ; full_smon :: StrongMonad M ; bind_as_join_sfmap `{Equiv A} `{Setoid B} `{!Setoid_Morphism (f : A → M B)} : bind f = join ∘ sfmap f }. math-classes-8.19.0/interfaces/naturals.v000066400000000000000000000041351460576051100203500ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.theory.categories MathClasses.varieties.semirings MathClasses.categories.varieties. Module bad. Class Naturals (A: semirings.Object) `{!InitialArrow A}: Prop := { naturals_initial:: Initial A }. End bad. Section initial_maps. Variable A: Type. Class NaturalsToSemiRing := naturals_to_semiring: ∀ B `{Mult B} `{Plus B} `{One B} `{Zero B}, A → B. Context `{NaturalsToSemiRing} `{SemiRing A} `{∀ `{SemiRing B}, SemiRing_Morphism (naturals_to_semiring B)}. Program Definition natural_initial_arrow: InitialArrow (semirings.object A) := λ y u, match u return A → y u with tt => naturals_to_semiring (y tt) end. Next Obligation. apply (@semirings.mor_from_sr_to_alg (λ _, A) _ _ (semirings.variety A)); apply _. Qed. Global Existing Instance natural_initial_arrow. (* For some reason if we try to make it an instance immediately upon definition, Program suddenly generates 5 subgoals.. *) Lemma natural_initial (same_morphism : ∀ `{SemiRing B} {h : A → B} `{!SemiRing_Morphism h}, naturals_to_semiring B = h) : Initial (semirings.object A). Proof. intros y [x h] [] ?. simpl in *. apply same_morphism. apply semirings.decode_variety_and_ops. apply (@semirings.decode_morphism_and_ops _ _ _ _ _ _ _ _ _ h). reflexivity. Qed. End initial_maps. #[global] Instance: Params (@naturals_to_semiring) 7 := {}. Class Naturals A {e plus mult zero one} `{U: NaturalsToSemiRing A} := { naturals_ring:: @SemiRing A e plus mult zero one ; naturals_to_semiring_mor:: ∀ `{SemiRing B}, SemiRing_Morphism (naturals_to_semiring A B) ; naturals_initial:: Initial (semirings.object A) }. (* Specializable operations: *) Class NatDistance N `{Equiv N} `{Plus N} := nat_distance_sig : ∀ x y : N, { z : N | x + z = y } + { z : N | y + z = x }. Definition nat_distance `{nd : NatDistance N} (x y : N) := match nat_distance_sig x y with | inl (n↾_) => n | inr (n↾_) => n end. #[global] Instance: Params (@nat_distance_sig) 4 := {}. #[global] Instance: Params (@nat_distance) 4 := {}. math-classes-8.19.0/interfaces/orders.v000066400000000000000000000211421460576051100200120ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra. (* In this file we describe interfaces for ordered structures. Since we are in a constructive setting we use a pseudo order instead of a total order. Therefore we also have to include an apartness relation. Obviously, in case we consider decidable structures these interfaces are quite inconvenient. Hence we will, later on, provide means to go back and forth between the usual classical notions and these constructive notions. On the one hand, if we have an ordinary (total) partial order (≤) with a corresponding strict order (<), we will prove that we can construct a FullPartialOrder and PseudoPartialOrder, respectively. On the other hand, if equality is decidable, we will prove that we have the usual properties like Trichotomy (<) and TotalRelation (≤). *) Class PartialOrder `{Ae : Equiv A} (Ale : Le A) : Prop := { po_setoid : Setoid A (* Making this a subclass makes instance search slow *) ; po_proper:: Proper ((=) ==> (=) ==> iff) (≤) ; po_preorder:: PreOrder (≤) ; po_antisym:: AntiSymmetric (≤) }. Class TotalOrder `{Ae : Equiv A} (Ale : Le A) : Prop := { total_order_po :: PartialOrder (≤) ; total_order_total :: TotalRelation (≤) }. (* We define a variant of the order theoretic definition of meet and join semilattices. Notice that we include a meet operation instead of the more common: ∀ x y, ∃ m, m ≤ x ∧ m ≤ y ∧ ∀ z, z ≤ x → z ≤ y → m ≤ z Our definition is both stronger and more convenient than the above. This is needed to prove equavalence with the algebraic definition. We do this in orders.lattices. *) Class MeetSemiLatticeOrder `{Ae : Equiv A} (Ale : Le A) `{Meet A} : Prop := { meet_sl_order :: PartialOrder (≤) ; meet_lb_l : ∀ x y, x ⊓ y ≤ x ; meet_lb_r : ∀ x y, x ⊓ y ≤ y ; meet_glb : ∀ x y z, z ≤ x → z ≤ y → z ≤ x ⊓ y }. Class JoinSemiLatticeOrder `{Ae : Equiv A} (Ale : Le A) `{Join A} : Prop := { join_sl_order :: PartialOrder (≤) ; join_ub_l : ∀ x y, x ≤ x ⊔ y ; join_ub_r : ∀ x y, y ≤ x ⊔ y ; join_lub : ∀ x y z, x ≤ z → y ≤ z → x ⊔ y ≤ z }. Class LatticeOrder `{Ae : Equiv A} (Ale : Le A) `{Meet A} `{Join A} : Prop := { lattice_order_meet :: MeetSemiLatticeOrder (≤) ; lattice_order_join :: JoinSemiLatticeOrder (≤) }. (* The strict order from the standard library does not include (=) and thus does not require (<) to be Proper. *) Class StrictSetoidOrder `{Ae : Equiv A} (Alt : Lt A) : Prop := { strict_setoid_order_setoid : Setoid A ; strict_setoid_order_proper :: Proper ((=) ==> (=) ==> iff) (<) ; strict_setoid_order_strict :: StrictOrder (<) }. (* The constructive notion of a total strict total order. Note that we get Proper (<) from cotransitivity. We will prove that (<) is in fact a StrictSetoidOrder. *) Class PseudoOrder `{Ae : Equiv A} `{Aap : Apart A} (Alt : Lt A) : Prop := { pseudo_order_setoid : StrongSetoid A ; pseudo_order_antisym : ∀ x y, ¬(x < y ∧ y < x) ; pseudo_order_cotrans :: CoTransitive (<) ; apart_iff_total_lt : ∀ x y, x ≶ y ↔ x < y ∨ y < x }. (* A partial order (≤) with a corresponding (<). We will prove that (<) is in fact a StrictSetoidOrder *) Class FullPartialOrder `{Ae : Equiv A} `{Aap : Apart A} (Ale : Le A) (Alt : Lt A) : Prop := { strict_po_setoid : StrongSetoid A ; strict_po_po :: PartialOrder (≤) ; strict_po_trans :: Transitive (<) ; lt_iff_le_apart : ∀ x y, x < y ↔ x ≤ y ∧ x ≶ y }. (* A pseudo order (<) with a corresponding (≤). We will prove that (≤) is in fact a PartialOrder. *) Class FullPseudoOrder `{Ae : Equiv A} `{Aap : Apart A} (Ale : Le A) (Alt : Lt A) : Prop := { full_pseudo_order_pseudo :: PseudoOrder Alt ; le_iff_not_lt_flip : ∀ x y, x ≤ y ↔ ¬y < x }. Section order_maps. Context {A B : Type} {Ae: Equiv A} {Ale: Le A} {Alt: Lt A} {Be: Equiv B} {Ble: Le B} {Blt: Lt B} (f : A → B). (* An Order_Morphism is just the factoring out of the common parts of OrderPreserving and OrderReflecting *) Class Order_Morphism := { order_morphism_mor : Setoid_Morphism f ; order_morphism_po_a : PartialOrder Ale ; order_morphism_po_b : PartialOrder Ble }. Class StrictOrder_Morphism := { strict_order_morphism_mor : Setoid_Morphism f ; strict_order_morphism_so_a : StrictSetoidOrder Alt ; strict_order_morphism_so_b : StrictSetoidOrder Blt }. Class OrderPreserving := { order_preserving_morphism :: Order_Morphism ; order_preserving : `(x ≤ y → f x ≤ f y) }. Class OrderReflecting := { order_reflecting_morphism :: Order_Morphism ; order_reflecting : `(f x ≤ f y → x ≤ y) }. Class OrderEmbedding := { order_embedding_preserving :: OrderPreserving ; order_embedding_reflecting :: OrderReflecting }. Class OrderIsomorphism `{!Inverse f} := { order_iso_embedding :: OrderEmbedding ; order_iso_surjective :: Surjective f }. Class StrictlyOrderPreserving := { strictly_order_preserving_morphism :: StrictOrder_Morphism ; strictly_order_preserving : `(x < y → f x < f y) }. Class StrictlyOrderReflecting := { strictly_order_reflecting_morphism :: StrictOrder_Morphism ; strictly_order_reflecting : `(f x < f y → x < y) }. Class StrictOrderEmbedding := { strict_order_embedding_preserving :: StrictlyOrderPreserving ; strict_order_embedding_reflecting :: StrictlyOrderReflecting }. End order_maps. #[global] Hint Extern 4 (?f _ ≤ ?f _) => apply (order_preserving f). #[global] Hint Extern 4 (?f _ < ?f _) => apply (strictly_order_preserving f). (* We define various classes to describe the order on the lower part of the algebraic hierarchy. This results in the notion of a PseudoSemiRingOrder, which specifies the order on the naturals, integers, rationals and reals. This notion is quite similar to a strictly linearly ordered unital commutative protoring in Davorin Lešnik's PhD thesis. *) Class SemiRingOrder `{Equiv A} `{Plus A} `{Mult A} `{Zero A} `{One A} (Ale : Le A) := { srorder_po :: PartialOrder Ale ; srorder_semiring : SemiRing A ; srorder_partial_minus : ∀ x y, x ≤ y → ∃ z, y = x + z ; srorder_plus :: ∀ z, OrderEmbedding (z +) ; nonneg_mult_compat : ∀ x y, PropHolds (0 ≤ x) → PropHolds (0 ≤ y) → PropHolds (0 ≤ x * y) }. Class StrictSemiRingOrder `{Equiv A} `{Plus A} `{Mult A} `{Zero A} `{One A} (Alt : Lt A) := { strict_srorder_so :: StrictSetoidOrder Alt ; strict_srorder_semiring : SemiRing A ; strict_srorder_partial_minus : ∀ x y, x < y → ∃ z, y = x + z ; strict_srorder_plus :: ∀ z, StrictOrderEmbedding (z +) ; pos_mult_compat : ∀ x y, PropHolds (0 < x) → PropHolds (0 < y) → PropHolds (0 < x * y) }. Class PseudoSemiRingOrder `{Equiv A} `{Apart A} `{Plus A} `{Mult A} `{Zero A} `{One A} (Alt : Lt A) := { pseudo_srorder_strict :: PseudoOrder Alt ; pseudo_srorder_semiring : SemiRing A ; pseudo_srorder_partial_minus : ∀ x y, ¬y < x → ∃ z, y = x + z ; pseudo_srorder_plus :: ∀ z, StrictOrderEmbedding (z +) ; pseudo_srorder_mult_ext :: StrongSetoid_BinaryMorphism (.*.) ; pseudo_srorder_pos_mult_compat : ∀ x y, PropHolds (0 < x) → PropHolds (0 < y) → PropHolds (0 < x * y) }. Class FullPseudoSemiRingOrder `{Equiv A} `{Apart A} `{Plus A} `{Mult A} `{Zero A} `{One A} (Ale : Le A) (Alt : Lt A) := { full_pseudo_srorder_pso :: PseudoSemiRingOrder Alt ; full_pseudo_srorder_le_iff_not_lt_flip : ∀ x y, x ≤ y ↔ ¬y < x }. (* Due to bug #2528 *) #[global] Hint Extern 7 (PropHolds (0 < _ * _)) => eapply @pos_mult_compat : typeclass_instances. #[global] Hint Extern 7 (PropHolds (0 ≤ _ * _)) => eapply @nonneg_mult_compat : typeclass_instances. (* Alternatively, we could have defined the standard notion of a RingOrder: Class RingOrder `{Equiv A} `{Plus A} `{Mult A} `{Zero A} (Ale : Le A) := { ringorder_po :: PartialOrder Ale ; ringorder_plus :: ∀ z, OrderPreserving (z +) ; ringorder_mult : ∀ x y, 0 ≤ x → 0 ≤ y → 0 ≤ x * y }. Unfortunately, this notion is too weak when we consider semirings (e.g. the naturals). Moreover, in case of rings, we prove that this notion is equivalent to our SemiRingOrder class (see orders.rings.from_ring_order). Hence we omit defining such a class. Similarly we prove that a FullSemiRingOrder and a FullPseudoRingOrder are equivalent. Class FullPseudoRingOrder `{Equiv A} `{Apart A} `{Plus A} `{Mult A} `{Zero A} (Ale : Le A) (Alt : Lt A) := { pseudo_ringorder_spo :: FullPseudoOrder Ale Alt ; pseudo_ringorder_mult_ext :: StrongSetoid_BinaryMorphism (.*.) ; pseudo_ringorder_plus :: ∀ z, StrictlyOrderPreserving (z +) ; pseudo_ringorder_mult : ∀ x y, 0 < x → 0 < y → 0 < x * y }. *) math-classes-8.19.0/interfaces/rationals.v000066400000000000000000000024631460576051100205150ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.integers MathClasses.implementations.field_of_fractions MathClasses.theory.integers. Section rationals_to_frac. Context (A : Type). Class RationalsToFrac := rationals_to_frac : ∀ B `{Integers B}, A → Frac B. End rationals_to_frac. (* We specify the Rationals as a field that contains the integers and can be embedded into the field of integers fractions. Since we do not want to fix a specific integer representation in this interface, we quantify over all integer implementations. However, when constructing an instance of the rationals it is generally inconvenient to prove that the required properties hold for all possible integer implementations. Therefore we provide a way (theory.rationals.alt_Build_Rationals) to construct a rationals implementation if the required properties hold for some specific implementation of the integers. *) Class Rationals A {e plus mult zero one neg recip} `{U : !RationalsToFrac A} : Prop := { rationals_field:: @DecField A e plus mult zero one neg recip ; rationals_frac :: ∀ `{Integers Z}, Injective (rationals_to_frac A Z) ; rationals_frac_mor :: ∀ `{Integers Z}, SemiRing_Morphism (rationals_to_frac A Z) ; rationals_embed_ints :: ∀ `{Integers Z}, Injective (integers_to_ring Z A) }. math-classes-8.19.0/interfaces/sequences.v000066400000000000000000000135131460576051100205120ustar00rootroot00000000000000Require Import Coq.Lists.List MathClasses.interfaces.abstract_algebra MathClasses.theory.categories MathClasses.theory.forget_algebra MathClasses.theory.forget_variety MathClasses.theory.rings MathClasses.interfaces.universal_algebra MathClasses.interfaces.functors MathClasses.categories.setoids MathClasses.categories.varieties. Require MathClasses.categories.product MathClasses.varieties.monoids MathClasses.categories.algebras MathClasses.categories.categories MathClasses.theory.setoids. Module ua := universal_algebra. #[global] Instance: Arrows Type := λ X Y, X → Y. (* todo: move elsewhere *) (* First, the nice clean high level encapsulated version: *) Class PoshSequence (free: setoids.Object → monoids.Object) `{!Fmap free} (inject: id ⇛ monoids.forget ∘ free) (extend: `((x ⟶ monoids.forget y) → (free x ⟶ y))): Prop := { sequence_adjunction: ηAdjunction _ _ inject extend ; extend_morphism: `(Setoid_Morphism (extend x y)) }. (* todo: how come extend_morphism isn't part of ηAdjunction? *) (* This looks very nice, but the encapsulation of the parameters makes it a bit awkward to work with. Hence, let us define a more down to earth version. *) Section practical. (* Here, again, are the ingredients, this time in somewhat more raw form: *) Class ExtendToSeq (free: Type → Type) := extend: ∀ {x y} `{!SgOp y} `{!MonUnit y}, (x → y) → (free x → y). (* todo: rename to extend_to_seq or something *) Class InjectToSeq (free: Type → Type) := inject: ∀ x, x → free x. (* todo: rename to singleton or something *) Context (free: Type → Type) {raw_fmap: Fmap free} `{∀ a, MonUnit (free a)} `{∀ a, SgOp (free a)} `{∀ a, Equiv a → Equiv (free a)} `{!InjectToSeq free} `{!ExtendToSeq free}. Class Sequence := { sequence_makes_monoids:: ∀ `{Setoid a}, Monoid (free a) ; sequence_inject_morphism:: ∀ `{Setoid a}, Setoid_Morphism (inject a) ; sequence_map_morphism:: ∀ `{Equiv x} `{Equiv y} (f: x → y), Setoid_Morphism f → Monoid_Morphism (raw_fmap _ _ f) ; sequence_fmap_proper: ∀ `{Equiv x} `{Equiv y} (f g: x → y), f = g → fmap free f = raw_fmap _ _ g ; sequence_fmap_id: ∀ `{Equiv x}, raw_fmap _ _ (@id x) = id ; sequence_fmap_comp: ∀ `{Equiv x} `{Equiv y} `{Equiv z} (f: y → z) (g: x → y), raw_fmap _ _ (f ∘ g) = raw_fmap _ _ f ∘ raw_fmap _ _ g ; sequence_extend_makes_morphisms:: ∀ `{Equiv x} `{Monoid y} (f: x → y), Setoid_Morphism f → Monoid_Morphism (extend f) ; sequence_inject_natural: ∀ `{Setoid A} `{Setoid B} (f: A → B), Setoid_Morphism f → inject B ∘ f = raw_fmap _ _ f ∘ inject A ; sequence_extend_commutes: ∀ `{Setoid x} `{Monoid y} (f: x → y), Setoid_Morphism f → extend f ∘ inject x = f ; sequence_only_extend_commutes: ∀ `{Setoid x} `{Monoid y} (f: x → y), Setoid_Morphism f → (∀ (g: free x → y), Monoid_Morphism g → g ∘ inject x = f → g = extend f) ; sequence_extend_morphism: ∀ `{Setoid x} `{Monoid y} (f g: x → y), Setoid_Morphism f → Setoid_Morphism g → f = g → extend f (free:=free) = extend g (free:=free) }. (* From this motley collection we can construct the encapsulated free/inject/extend: *) Context `{PS: Sequence}. Program Definition posh_free (X: setoids.Object): monoids.Object := monoids.object (free X). Program Instance posh_fmap: functors.Fmap posh_free := λ _ _ X _, raw_fmap _ _ X. Next Obligation. apply monoids.encode_morphism_only. destruct X. apply _. Qed. Instance: Functor posh_free posh_fmap. Proof with try apply _. constructor... repeat intro. constructor... repeat intro. simpl. apply sequence_fmap_proper. intro. apply H2. reflexivity. repeat intro. simpl. apply sequence_fmap_id. reflexivity. repeat intro. simpl. apply sequence_fmap_comp. reflexivity. Qed. Program Definition posh_inject: id ⇛ monoids.forget ∘ posh_free := λ a, inject a. (* Needed for some type conversions. *) Typeclasses Transparent compose. Program Definition posh_extend (x: setoids.Object) (y: monoids.Object) (X: x ⟶ monoids.forget y): posh_free x ⟶ y := λ u, match u return posh_free x u → y u with tt => extend X end. Next Obligation. apply monoids.encode_morphism_only. destruct X. simpl in *. apply (sequence_extend_makes_morphisms _). apply _. Qed. (* ... and show that they form a posh sequence: *) Instance: NaturalTransformation posh_inject. Proof. unfold NaturalTransformation. intros [???] [???] [??] ?? E. simpl in *. rewrite E. apply sequence_inject_natural. apply _. reflexivity. Qed. Goal @PoshSequence posh_free posh_fmap posh_inject posh_extend. Proof. constructor. constructor; try apply _. intros [x xE xH] y [f fM]. pose proof (@monoids.decode_variety_and_ops y _ _ _). split. repeat intro. simpl in *. rewrite H3. symmetry. apply (@sequence_extend_commutes PS x _ _ _ _ _ _ H2 f fM). reflexivity. unfold compose. intros [x0 h] H4 [] a. unfold equiv, setoids.Equiv_instance_0 in H4. simpl in *. apply (@sequence_only_extend_commutes PS x _ _ _ _ _ _ H2 f _ (x0 tt)). apply (@monoids.decode_morphism_and_ops _ _ _ _ _ _ _ _ _ h). intros. symmetry. apply H4. reflexivity. unfold posh_extend. intros [x ??] [y ?? yV]. constructor; try apply _. intros [] [] E [] a. simpl in *. apply (@sequence_extend_morphism PS x _ _ _ _ _ _ (@monoids.decode_variety_and_ops _ _ _ yV) _ _ _ _). intro. apply E. reflexivity. Qed. (* todo: clean up *) Definition fold `{MonUnit M} `{SgOp M}: free M → M := extend id. Global Instance fold_mor `{Monoid M}: Monoid_Morphism (fold (M:=M)). Proof. apply _. Qed. End practical. math-classes-8.19.0/interfaces/ua_basic.v000066400000000000000000000070151460576051100202650ustar00rootroot00000000000000Require MathClasses.implementations.ne_list. Require Import Coq.Lists.List MathClasses.interfaces.abstract_algebra. Local Notation ne_list := ne_list.L. Section with_sorts. Variable Sorts: Set. (* For single-sorted algebras, Sorts will typically be unit. *) (* OpType describes the type of an operation in an algebra. Note that higher order function types are excluded: *) Definition OpType := ne_list Sorts. Definition result: OpType → Sorts := @ne_list.last _. Variable carrier: Sorts → Type. (* Given a Type for each sort, we can map the operation type descriptions to real function types: *) Fixpoint op_type (o: OpType): Type := match o with | ne_list.one a => carrier a | ne_list.cons a g => carrier a → op_type g end. (* This is just: Definition op_type: OpType → Type := ne_list.foldr1 (→) ∘ ne_list.map carrier. Unfortunately, in that formulation [simpl] never reduces it, which is extremely annoying... *) (* We use extensional equivalence for such generated function types: *) Context `{e: ∀ s, Equiv (carrier s)}. Fixpoint op_type_equiv o: Equiv (op_type o) := match o with | ne_list.one _ => _: Equiv (carrier _) (*e A*) | ne_list.cons A g => (e A ==> op_type_equiv g)%signature end. Existing Instance op_type_equiv. (* There's no [Global Instance Fixpoint]. *) Global Instance sig_type_sym `{∀ s, Symmetric (e s)} {o} : Symmetric (op_type_equiv o). Proof. induction o; simpl; firstorder. Qed. (* We need either reflexivity or symmetry of e in order to get transitivity of op_type_equiv: *) Global Instance sig_type_trans `{∀ s, Reflexive (e s)} `{∀ s, Transitive (e s)} {o}: Transitive (op_type_equiv o). Proof. induction o; simpl. firstorder. intros ? y ???? y0 ?. transitivity (y y0); firstorder. Qed. Hint Unfold op_type. Global Instance sig_type_trans' `{∀ s, Symmetric (e s)} `{∀ s, Transitive (e s)} {o}: Transitive (op_type_equiv o). Proof with auto. induction o; simpl... intros x y ? ? H2 x0 y0 ?. transitivity (y y0)... apply H2. transitivity x0; firstorder. Qed. (* This is the closest i've been able to get to reflexivity thus far: *) Lemma sig_type_refl `{∀ a, Reflexive (e a)} (o: OpType) a (x: op_type (ne_list.cons a o)) y: Proper (=) x → op_type_equiv o (x y) (x y). Proof. intro H0. apply H0. reflexivity. Qed. (* Lemma sig_type_refl' (o: OpType) a (x: op_type (function a o)): Proper (=) x → op_type_equiv _ x x. Proof. intro H0. apply H0. Qed. *) End with_sorts. Arguments op_type {Sorts} _ _. (* Avoid eager application *) #[global] Hint Extern 0 (Equiv (op_type _ _ )) => eapply @op_type_equiv : typeclass_instances. Inductive Signature: Type := { sorts: Set ; operation:> Set ; operation_type:> operation → OpType sorts }. Definition single_sorted_signature {Op: Set} (arities: Op → nat): Signature := Build_Signature unit Op (ne_list.replicate_Sn tt ∘ arities). (* An implementation of a signature for a given realization of the sorts is simply a function (of the right type) for each operation: *) Class AlgebraOps (σ: Signature) (A: sorts σ → Type) := algebra_op: ∀ o, op_type A (σ o). (* .. which, if they are proper with respect to a bona fide setoid equality, form an algebra: *) Class Algebra (σ: Signature) (carriers: sorts σ → Type) {e: ∀ a, Equiv (carriers a)} `{AlgebraOps σ carriers}: Prop := { algebra_setoids:: ∀ a, Setoid (carriers a) ; algebra_propers:: ∀ o: σ, Proper (=) (algebra_op o) }. math-classes-8.19.0/interfaces/universal_algebra.v000066400000000000000000000230051460576051100222010ustar00rootroot00000000000000Require MathClasses.theory.setoids MathClasses.implementations.ne_list. Require Import Coq.Lists.List MathClasses.interfaces.abstract_algebra MathClasses.misc.util MathClasses.theory.jections. Require Export MathClasses.interfaces.ua_basic. Section for_signature. Variable σ: Signature. Notation OpType := (OpType (sorts σ)). Inductive Term (V: Type): OpType → Type := | Var: V → ∀ a, Term V (ne_list.one a) | App t y: Term V (ne_list.cons y t) → Term V (ne_list.one y) → Term V t | Op o: Term V (σ o). Arguments Var {V} _ _. Fixpoint map_var `(f: V → W) `(t: Term V o): Term W o := match t in Term _ o return Term W o with | Var v s => Var (f v) s | App _ _ _ x y => App _ _ _ (map_var f x) (map_var f y) | Op _ s => Op _ s end. (* Term has OpType as an index, which means we can have terms with function types (no surprise there). However, often we want to prove properties that only speak of nullary terms: *) Definition Term0 v sort := Term v (ne_list.one sort). Section applications_ind. Context V (P: ∀ {a}, Term0 V a → Type). Arguments P {a} _. (* Proving such properties for nullary terms directly using Term's induction principle is problematic because it requires a property over terms of /any/ arity. Hence, we must first transform P into a statement about terms of all arities. Roughly speaking, we do this by saying [∀ x0...xN, P (App (... (App f x0) ...) xN)] for a term f of arity N. *) Fixpoint applications {ot}: Term V ot → Type := match ot with | ne_list.one x => @P x | ne_list.cons x y => λ z, ∀ v, P v → applications (App V _ _ z v) end. (* To prove P/applications by induction, we can then use: *) Lemma applications_rect: (∀ v a, P (Var v a)) → (∀ o, applications (Op _ o)) → (∀ a (t: Term0 V a), P t). Proof. intros X0 X1 ??. cut (applications t). intros. assumption. induction t; simpl. apply X0. apply IHt1; exact IHt2. apply X1. Defined. (* todo: write as term *) End applications_ind. (* We parameterized Term over the index set for variables, because we will wants terms /with/ variables when speaking of identities in equational theories, but will want terms /without/ variables when constructing initial objects in variety categories generically in theory/ua_initial (which we get by taking False as the variable index set). *) Definition T := Term nat. Definition T0 := Term0 nat. Definition Identity t := prod (T t) (T t). Definition Identity0 sort := Identity (ne_list.one sort). (* While Identity0 is the one we usually have in mind, the generalized version for arbitrary op_types is required to make induction proofs work. *) Definition mkIdentity0 {sort}: T (ne_list.one sort) → T (ne_list.one sort) → Identity0 sort := pair. (* The laws in an equational theory will be entailments of identities for any of the sorts: *) Record Entailment (P: Type): Type := { entailment_premises: list P; entailment_conclusion: P }. Definition EqEntailment := Entailment (sigT Identity0). (* We also introduce a more general class of statements that we use to conveniently transfer statements between different models of a variety: *) Inductive Statement: Type := | Eq t (i: Identity t) | Impl (a b: Statement) | Conj (a b: Statement) | Disj (a b: Statement) | Ext (P: Prop). (* Statements are a strict generalization of EqEntailments. We cannot use the former for the laws, though, because they are too liberal (i.e. not equational) to support product varieties. We do have two injections: *) Definition identity_as_eq (s: sigT Identity0): Statement := Eq _ (projT2 s). Coercion identity_as_entailment sort (e: Identity0 sort): EqEntailment := Build_Entailment _ nil (existT _ _ e). Coercion entailment_as_statement (e: EqEntailment): Statement := (fold_right Impl (identity_as_eq (entailment_conclusion _ e)) (map identity_as_eq (entailment_premises _ e))). Definition entailment_as_conjunctive_statement (e: EqEntailment): Statement := Impl (fold_right Conj (Ext True) (map identity_as_eq (entailment_premises _ e))) (identity_as_eq (entailment_conclusion _ e)). (* The first one (the coercion) converts an entailment of the form (A, B, C |- D) into a statement of the form (A → B → C → D), while the second converts the same entailment into a statement of the form ((A ∧ B ∧ C) → D). Both have their uses in induction proofs. *) (* To be able to state that laws hold in a model of a variety, we must be able to evaluate terms using the model's implementation and using arbitrary variable assignments: *) Section Vars. Context (A: sorts σ → Type) (V: Type) `{e: ∀ a, Equiv (A a)} `{∀ a, Equivalence (e a)}. Definition Vars := ∀ a, V → A a. Global Instance ua_vars_equiv: Equiv Vars := @pointwise_dependent_relation (sorts σ) (λ a, V → A a) (λ _, pointwise_relation _ (=)). Global Instance: Equivalence ((=): relation Vars) := {}. End Vars. Definition no_vars x: Vars x False := λ _, False_rect _. (* Given an assignment mapping variables to closed terms, we can close open terms: *) Fixpoint close {V} {o} (v: Vars (λ x, Term False (ne_list.one x)) V) (t: Term V o): Term False o := match t in Term _ o return Term False o with | Var x y => v y x | App _ x y z r => App _ x y (close v z) (close v r) | Op _ o => Op _ o end. Section eval. Context `{Algebra σ A}. Fixpoint eval {V} {n: OpType} (vars: Vars A V) (t: Term V n) {struct t}: op_type A n := match t with | Var v a => vars a v | Op _ o => algebra_op o | App _ n a f p => eval vars f (eval vars p) end. Global Instance eval_proper {V} (n: OpType): Proper ((=) ==> eq ==> (=)) (@eval V n). Proof with auto. intros x y E a _ []. induction a. apply E... apply IHa1... simpl. apply algebra_propers. Qed. Global Instance eval_strong_proper {V} (n: OpType): Proper ((pointwise_dependent_relation (sorts σ) _ (λ _, pointwise_relation V eq)) ==> eq ==> eq) (@eval V n). Proof with auto. intros x y E a _ []. unfold pointwise_dependent_relation in E. unfold pointwise_relation in E. induction a; simpl. apply E... congruence. reflexivity. Qed. Hint Extern 4 (Equiv (Term _ _)) => exact eq: typeclass_instances. Hint Extern 4 (Equiv (Term0 _ _)) => exact eq: typeclass_instances. Instance: ∀ V n v, Setoid_Morphism (@eval V (ne_list.one n) v). Proof. constructor; try apply _. unfold Setoid. apply _. destruct H0. apply _. Qed. Fixpoint app_tree {V} {o}: Term V o → op_type (Term0 V) o := match o with | ne_list.one _ => id | ne_list.cons _ _ => λ x y, app_tree (App _ _ _ x y) end. (* Instance: AlgebraOps σ (Term0 V) := λ _ x => app_tree (Op _ x). (* todo: these two are now duplicate with open_terms *) Instance: Algebra σ (Term0 V). Proof. constructor. intro. unfold Setoid. apply _. intro. change (Proper (=) (app_tree (Op V o))). generalize (Op V o). induction (operation_type σ o). reflexivity. simpl. repeat intro. subst. apply IHo0. Qed. *) Lemma eval_map_var `(f: V → W) v s (t: Term V s): eval v (map_var f t) ≡ eval (λ s, v s ∘ f) t. Proof. induction t; simpl. reflexivity. congruence. reflexivity. Qed. Definition eval_stmt (vars: Vars A nat): Statement → Prop := fix F (s: Statement) := match s with | Eq _ i => eval vars (fst i) = eval vars (snd i) | Impl a b => F a → F b | Ext P => P | Conj a b => F a ∧ F b | Disj a b => F a ∨ F b end. Global Instance eval_stmt_proper: Proper ((=) ==> eq ==> iff) eval_stmt. Proof with auto. intros v v' ve s s' se. subst. induction s'; simpl; try solve [intuition]. split; intros E. transitivity (eval v (fst i)). apply eval_proper... symmetry... transitivity (eval v (snd i))... apply eval_proper... transitivity (eval v' (fst i)). apply eval_proper... rewrite E. apply eval_proper... Qed. Definition boring_eval_entailment (vars: Vars A nat) (ee: EqEntailment): eval_stmt vars ee ↔ eval_stmt vars (entailment_as_conjunctive_statement ee). Proof. destruct ee. simpl. induction entailment_premises0; simpl; intuition. Qed. End eval. End for_signature. (* Avoid eager application *) #[global] Remove Hints ua_vars_equiv : typeclass_instances. #[global] Hint Extern 0 (Equiv (Vars _ _ _)) => eapply @ua_vars_equiv : typeclass_instances. #[global] Instance: Params (@eval_stmt) 3 := {}. (* And with that, we define equational theories and varieties: *) Record EquationalTheory := { et_sig:> Signature ; et_laws:> EqEntailment et_sig → Prop }. Class InVariety (et: EquationalTheory) (carriers: sorts et → Type) {e: ∀ a, Equiv (carriers a)} `{!AlgebraOps et carriers}: Prop := { variety_algebra:: Algebra et carriers ; variety_laws: ∀ s, et_laws et s → ∀ vars, eval_stmt et vars s }. Module op_type_notations. Global Infix "-=>" := (ne_list.cons) (at level 95, right associativity). End op_type_notations. (* todo: get rid of *) Module notations. Global Infix "===" := (mkIdentity0 _) (at level 70, no associativity). Global Infix "-=>" := (Impl _) (at level 95, right associativity). End notations. math-classes-8.19.0/interfaces/vectorspace.v000066400000000000000000000124241460576051100210350ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders. (** Scalar multiplication function class *) Class ScalarMult K V := scalar_mult: K → V → V. #[global] Instance: Params (@scalar_mult) 3 := {}. Infix "·" := scalar_mult (at level 50) : mc_scope. Notation "(·)" := scalar_mult (only parsing) : mc_scope. Notation "( x ·)" := (scalar_mult x) (only parsing) : mc_scope. Notation "(· x )" := (λ y, y · x) (only parsing) : mc_scope. (** The inproduct function class *) Class Inproduct K V := inprod : V → V → K. #[global] Instance: Params (@inprod) 3 := {}. Notation "(⟨⟩)" := (inprod) (only parsing) : mc_scope. Notation "⟨ u , v ⟩" := (inprod u v) (at level 51) : mc_scope. Notation "⟨ u , ⟩" := (λ v, ⟨u,v⟩) (at level 50, only parsing) : mc_scope. Notation "⟨ , v ⟩" := (λ u, ⟨u,v⟩) (at level 50, only parsing) : mc_scope. Notation "x ⊥ y" := (⟨x,y⟩ = 0) (at level 70) : mc_scope. (** The norm function class *) Class Norm K V := norm : V → K. #[global] Instance: Params (@norm) 2 := {}. Notation "∥ L ∥" := (norm L) (at level 50) : mc_scope. Notation "∥·∥" := norm (only parsing) : mc_scope. (** Let [M] be an R-Module. *) Class Module (R M : Type) {Re Rplus Rmult Rzero Rone Rnegate} {Me Mop Munit Mnegate} {sm : ScalarMult R M} := { lm_ring :: @Ring R Re Rplus Rmult Rzero Rone Rnegate ; lm_group :: @AbGroup M Me Mop Munit Mnegate ; lm_distr_l :: LeftHeteroDistribute (·) (&) (&) ; lm_distr_r :: RightHeteroDistribute (·) (+) (&) ; lm_assoc :: HeteroAssociative (·) (·) (·) (.*.) ; lm_identity :: LeftIdentity (·) 1 ; scalar_mult_proper :: Proper ((=) ==> (=) ==> (=)) sm }. (* TODO K is commutative, so derive right module laws? *) (** A module with a seminorm. *) Class Seminormed {R Re Rplus Rmult Rzero Rone Rnegate Rle Rlt Rapart} {M Me Mop Munit Mnegate Smult} `{!Abs R} (n : Norm R M) := (* We have a module *) { snm_module :: @Module R M Re Rplus Rmult Rzero Rone Rnegate Me Mop Munit Mnegate Smult ; snm_order :: @FullPseudoSemiRingOrder R Re Rapart Rplus Rmult Rzero Rone Rle Rlt (* With respect to which our norm preserves the following: *) ; snm_scale : ∀ a v, ∥a · v∥ = (abs a) * ∥v∥ (* positive homgeneity *) ; snm_triangle : ∀ u v, ∥u & v∥ ≤ ∥u∥ + ∥v∥ (* triangle inequality *) }. (** [K] is the field of scalars, [V] the abelian group of vectors, and together with a scalar multiplication operation, they satisfy the Module laws. *) Class VectorSpace (K V : Type) {Ke Kplus Kmult Kzero Kone Knegate Krecip} (* scalar operations *) {Ve Vop Vunit Vnegate} (* vector operations *) {sm : ScalarMult K V} := { vs_field :: @DecField K Ke Kplus Kmult Kzero Kone Knegate Krecip ; vs_abgroup :: @AbGroup V Ve Vop Vunit Vnegate ; vs_module :: @Module K V Ke Kplus Kmult Kzero Kone Knegate Ve Vop Vunit Vnegate sm }. (** Given some vector space V over a ordered field K, we define the inner product space *) Class InnerProductSpace (K V : Type) {Ke Kplus Kmult Kzero Kone Knegate Krecip} (* scalar operations *) {Ve Vop Vunit Vnegate} (* vector operations *) {sm : ScalarMult K V} {inp: Inproduct K V} {Kle: Le K} := { in_vectorspace :: @VectorSpace K V Ke Kplus Kmult Kzero Kone Knegate Krecip Ve Vop Vunit Vnegate sm ; in_srorder :: SemiRingOrder Kle ; in_comm :: Commutative inprod ; in_linear_l : ∀ a u v, ⟨a·u,v⟩ = a*⟨u,v⟩ ; in_nonneg :: ∀ v, PropHolds (0 ≤ ⟨v,v⟩) (* TODO Le to strong? *) ; in_mon_unit_zero :: ∀ v, 0 = ⟨v,v⟩ <-> v = mon_unit ; inprod_proper :: Proper ((=) ==> (=) ==> (=)) (⟨⟩) }. (* TODO complex conjugate? Section proof_in_linear_r. Context `{InnerProductSpace}. Lemma in_linear_r a u v : ⟨u,a·v⟩ = a*⟨u,v⟩. Proof. rewrite !(commutativity u). apply in_linear_l. Qed. End proof_in_linear_r. *) (* This is probably a bad idea? (because ∣ ≠ |) Notation "∣ a ∣" := (abs a). *) Class SemiNormedSpace (K V : Type) `{a:Abs K} `{n : @Norm K V} (* scalar and vector norms *) {Ke Kplus Kmult Kzero Kone Knegate Krecip} (* scalar operations *) {Ve Vop Vunit Vnegate} (* vector operations *) {sm : ScalarMult K V} := { sn_vectorspace :: @VectorSpace K V Ke Kplus Kmult Kzero Kone Knegate Krecip Ve Vop Vunit Vnegate sm ; sn_nonneg : ∀ v, 0 ≤ ∥v∥ (* non-negativity *) ; sn_scale : ∀ a v, ∥a · v∥ = (abs a) * ∥v∥ (* positive homgeneity *) ; sn_triangle : ∀ u v, ∥u & v∥ ≤ ∥u∥ + ∥v∥ (* triangle inequality *) }. (* For normed spaces: n_separates : ∀ (v:V), ∥v∥ = 0 ↔ v = unit *) (* - Induced metric: d x y := ∥ x - y ∥ - Induced norm. If the metric is 1). translation invariant: d x y = d (x + a) (y + a) 2). and homogeneous: d (α * x) (α * y) = ∣α∣ * (d x y) then we can define the norm as: ∥ x ∥ := d 0 x - Same for seminorm *) math-classes-8.19.0/meta.yml000066400000000000000000000062241460576051100156570ustar00rootroot00000000000000--- fullname: Math Classes shortname: math-classes organization: coq-community community: true action: true ci_extra_dev: true doi: 10.1017/S0960129511000119 synopsis: >- A library of abstract interfaces for mathematical structures in Coq description: | Math classes is a library of abstract interfaces for mathematical structures, such as: * Algebraic hierarchy (groups, rings, fields, …) * Relations, orders, … * Categories, functors, universal algebra, … * Numbers: N, Z, Q, … * Operations, (shift, power, abs, …) It is heavily based on Coq’s new type classes in order to provide: structure inference, multiple inheritance/sharing, convenient algebraic manipulation (e.g. rewriting) and idiomatic use of notations. publications: - pub_doi: 10.1017/S0960129511000119 pub_url: https://arxiv.org/abs/1102.1323 pub_title: Type Classes for Mathematics in Type Theory authors: - name: Eelis van der Weegen initial: true - name: Bas Spitters initial: true - name: Robbert Krebbers initial: true maintainers: - name: Bas Spitters nickname: spitters opam-file-maintainer: b.a.w.spitters@gmail.com license: fullname: MIT License identifier: MIT supported_coq_versions: text: Coq 8.18 or later (use releases for other Coq versions) opam: '{(>= "8.18" & < "8.20~") | (= "dev")}' tested_coq_opam_versions: - version: dev - version: "8.19" - version: "8.18" dependencies: - opam: name: coq-bignums nix: bignums description: "[BigNums](https://github.com/coq/bignums)" namespace: MathClasses build: | ## Building and installation instructions The easiest way to install the latest released version of Math Classes is via [OPAM](https://opam.ocaml.org/doc/Install.html): ```shell opam repo add coq-released https://coq.inria.fr/opam/released opam install coq-math-classes ``` To instead build and install manually, do: ``` shell git clone https://github.com/coq-community/math-classes.git cd math-classes ./configure.sh make # or make -j make install ``` documentation: | ## Directory structure ### categories/ Proofs that certain structures form categories. ### functors/ ### interfaces/ Definitions of abstract interfaces/structures. ### implementations/ Definitions of concrete data structures and algorithms, and proofs that they are instances of certain structures (i.e. implement certain interfaces). ### misc/ Miscellaneous things. ### orders/ Theory about orders on different structures. ### quote/ Prototype implementation of type class based quoting. To be integrated. ### theory/ Proofs of properties of structures. ### varieties/ Proofs that certain structures are varieties, and translation to/from type classes dedicated to these structures (defined in interfaces/). The reason we treat categories and varieties differently from other structures (like groups and rings) is that they are like meta-interfaces whose implementations are not concrete data structures and algorithms but are themselves abstract structures. To be able to distinguish the various arrows, we recommend using a variable width font. --- math-classes-8.19.0/misc/000077500000000000000000000000001460576051100151355ustar00rootroot00000000000000math-classes-8.19.0/misc/JMrelation.v000066400000000000000000000021331460576051100173670ustar00rootroot00000000000000(* JMeq without the [eq] hard-wiring. Meant for use with [Require] only, not [Import]. *) Require Import Relation_Definitions Setoid. Require Export Unicode.Utf8 Utf8_core. Inductive R {A: Type} (r: relation A) (x: A): forall B: Type, relation B → B → Prop := relate y: r x y → R r x A r y. Lemma reflexive A (x:A) (Ra: relation A) `{!Reflexive Ra}: R Ra x _ Ra x. Proof. apply relate. reflexivity. Qed. Lemma symmetric A B (x:A) (Ra: relation A) (Rb: relation B) (y:B) `{!Symmetric Ra}: R Ra x _ Rb y → R Rb y _ Ra x. Proof. destruct 1. apply relate. symmetry. assumption. Qed. Lemma transitive A B C (Ra: relation A) (Rb: relation B) (Rc: relation C) (a:A) (b:B) (c:C) `{!Transitive Ra}: R Ra a _ Rb b → R Rb b _ Rc c → R Ra a _ Rc c. Proof. destruct 1. destruct 1. apply relate. transitivity y; assumption. Qed. Require Import Coq.Logic.Eqdep. Lemma unJM A (r: relation A) (x y:A) r' (E: R r x A r' y): r x y. Proof. simple inversion E. rewrite (eq_rect_eq__eq_dep_eq _ (eq_rect_eq _) _ _ _ _ (eq_sigT_eq_dep _ _ _ _ _ _ H2)). firstorder. Qed. (* Warning: Depends on proof_irrelevance. *) math-classes-8.19.0/misc/benchmarks_nobuild.v000066400000000000000000000075231460576051100211640ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.integers MathClasses.interfaces.additional_operations MathClasses.implementations.dyadics MathClasses.implementations.fast_integers. Section wolfram_sqrt. Context `{Integers Z} `{!RingOrder oZ} `{!TotalOrder oZ} `{precedes_dec : ∀ (x y : Z), Decision (x ≤ y)} `{!NatPowSpec Z (Z⁺) pw} `{!ShiftLSpec Z (Z⁺) sl}. Fixpoint root_loop (x : Dyadic Z) (n : nat) : Dyadic Z * Dyadic Z := match n with | O => (x, 0) | S n => let (r, s) := root_loop x n in if decide_rel (≤) (s + 1) r then ((r - (s + 1)) ≪ 2, (s + 2) ≪ 1) else (r ≪ 2, s ≪ 1) end. End wolfram_sqrt. Definition fast_root_loop := root_loop (Z:=fastZ). Let n : nat := Eval vm_compute in (10 * 10 * 10 * 10)%nat. Time Eval vm_compute in (snd (fast_root_loop 2 n)). Time Eval vm_compute in ( (fun _ _ _ _ _ _ _ _ _ _ => true) (snd (fast_root_loop 2 n)) (snd (fast_root_loop 2 n)) (snd (fast_root_loop 2 n)) (snd (fast_root_loop 2 n)) (snd (fast_root_loop 2 n)) (snd (fast_root_loop 2 n)) (snd (fast_root_loop 2 n)) (snd (fast_root_loop 2 n)) (snd (fast_root_loop 2 n)) (snd (fast_root_loop 2 n))). Require Import BigZ. Open Scope bigZ_scope. Notation bigD := (Dyadic bigZ). Definition BigD_0 : bigD := (0 ▼ 0). Definition BigD_1 : bigD := (1 ▼ 0). Definition BigD_2 : bigD := (2 ▼ 0). Definition BigD_plus (x y : bigD) : bigD := match BigZ.compare (expo x) (expo y) with | Gt => BigZ.shiftl (mant x) (expo x - expo y) + mant y ▼ BigZ.min (expo x) (expo y) | _ => mant x + BigZ.shiftl (mant y) (expo y - expo x) ▼ BigZ.min (expo x) (expo y) end. Definition BigD_opp (x : bigD) : bigD := -mant x ▼ expo x. Definition BigD_mult (x y : bigD) : bigD := mant x * mant y ▼ expo x + expo y. Definition BigD_shiftl (x : bigD) (n : bigZ) : bigD := mant x ▼ expo x + n. Definition BigD_compare (x y : bigD) : comparison := match BigZ.compare (expo x) (expo y) with | Gt => BigZ.compare (BigZ.shiftl (mant x) (expo x - expo y)) (mant y) | _ => BigZ.compare (mant x) (BigZ.shiftl (mant y) (expo y - expo x)) end. Fixpoint root_loop_alt (x : bigD) (n : nat) : bigD * bigD := match n with | O => (x, BigD_0) | S n => let (r, s) := root_loop_alt x n in match BigD_compare (BigD_plus s BigD_1) r with | Gt => (BigD_shiftl r 2, BigD_shiftl s 1) | _ => (BigD_shiftl (BigD_plus r (BigD_opp (BigD_plus s BigD_1))) 2, BigD_shiftl (BigD_plus s BigD_2) 1) end end. Time Eval vm_compute in (mant (snd (root_loop_alt BigD_2 n))). Time Eval vm_compute in ( (fun _ _ _ _ _ _ _ _ _ _ => true) (snd (root_loop_alt BigD_2 n)) (snd (root_loop_alt BigD_2 n)) (snd (root_loop_alt BigD_2 n)) (snd (root_loop_alt BigD_2 n)) (snd (root_loop_alt BigD_2 n)) (snd (root_loop_alt BigD_2 n)) (snd (root_loop_alt BigD_2 n)) (snd (root_loop_alt BigD_2 n)) (snd (root_loop_alt BigD_2 n)) (snd (root_loop_alt BigD_2 n))). Definition BigD_4 : bigD := (4 ▼ 0). Fixpoint root_loop_alt_mult (x : bigD) (n : nat) : bigD * bigD := match n with | O => (x, BigD_0) | S n => let (r, s) := root_loop_alt_mult x n in match BigD_compare (BigD_plus s BigD_1) r with | Gt => (BigD_mult BigD_4 r, BigD_mult BigD_2 s) | _ => (BigD_mult BigD_4 (BigD_plus r (BigD_opp (BigD_plus s BigD_1))), BigD_mult BigD_2 (BigD_plus s BigD_2)) end end. Time Eval vm_compute in (mant (snd (root_loop_alt_mult BigD_2 n))). Time Eval vm_compute in ( (fun _ _ _ _ _ _ _ _ _ _ => true) (snd (root_loop_alt_mult BigD_2 n)) (snd (root_loop_alt_mult BigD_2 n)) (snd (root_loop_alt_mult BigD_2 n)) (snd (root_loop_alt_mult BigD_2 n)) (snd (root_loop_alt_mult BigD_2 n)) (snd (root_loop_alt_mult BigD_2 n)) (snd (root_loop_alt_mult BigD_2 n)) (snd (root_loop_alt_mult BigD_2 n)) (snd (root_loop_alt_mult BigD_2 n)) (snd (root_loop_alt_mult BigD_2 n))). math-classes-8.19.0/misc/decision.v000066400000000000000000000111451460576051100171230ustar00rootroot00000000000000Require Import MathClasses.interfaces.canonical_names MathClasses.misc.util. Require Bool. Class Decision P := decide: sumbool P (¬P). Arguments decide _ {Decision}. #[global] Instance: ∀ P, Decision P → Stable P. Proof. firstorder. Qed. Ltac case_decide := match goal with | H : context [@decide ?P ?dec] |- _ => case (@decide P dec) in * | |- context [@decide ?P ?dec] => case (@decide P dec) in * end. Ltac solve_trivial_decision := match goal with | [ |- Decision (?P) ] => apply _ | [ |- sumbool ?P (¬?P) ] => change (Decision P); apply _ end. Ltac solve_decision := first [solve_trivial_decision | unfold Decision; decide equality; solve_trivial_decision]. (* We cannot state this as Proper (iff ==> iffT) Decision due to the lack of setoid rewriting in Type *) Program Definition decision_proper (P Q : Prop) (PiffQ : P ↔ Q) (P_dec : Decision P) : Decision Q := match P_dec with | left _ => left _ | right _ => right _ end. Solve Obligations with (program_simpl; tauto). Definition bool_decide (P : Prop) `{dec : !Decision P} : bool := if dec then true else false. Lemma bool_decide_true `{dec : Decision P} : bool_decide P ≡ true ↔ P. Proof. unfold bool_decide. split; intro; destruct dec; firstorder with bool. Qed. Lemma bool_decide_false `{dec : !Decision P} : bool_decide P ≡ false ↔ ¬P. Proof. unfold bool_decide. split; intro; destruct dec; firstorder with bool. Qed. (* Because [vm_compute] evaluates terms in [Prop] eagerly and does not remove dead code we need the decide_rel hack. Suppose we have [(x = y) =def (f x = f y)], now: bool_decide (x = y) → bool_decide (f x = f y) → ... As we see, the dead code [f x] and [f y] is actually evaluated, which is of course an utter waste. Therefore we introduce decide_rel and bool_decide_rel. bool_decide_rel (=) x y → bool_decide_rel (λ a b, f a = f b) x y → ... Now the definition of equality remains under a lambda and our problem does not occur anymore! *) Definition decide_rel `(R : A → B → Prop) {dec : ∀ x y, Decision (R x y)} (x : A) (y : B) : Decision (R x y) := dec x y. Definition bool_decide_rel `(R : relation A) {dec : ∀ x y, Decision (R x y)} : A → A → bool := λ x y, if dec x y then true else false. Lemma bool_decide_rel_true `(R : relation A) {dec : ∀ x y, Decision (R x y)} : ∀ x y, bool_decide_rel R x y ≡ true ↔ R x y. Proof. unfold bool_decide_rel. split; intro; destruct dec; firstorder with bool. Qed. Lemma bool_decide_rel_false `(R : relation A)`{dec : ∀ x y, Decision (R x y)} : ∀ x y, bool_decide_rel R x y ≡ false ↔ ¬R x y. Proof. unfold bool_decide_rel. split; intro; destruct dec; firstorder with bool. Qed. Program Definition decision_from_bool_decide {P b} (prf : b ≡ true ↔ P) : Decision P := match b with true => left _ | false => right _ end. Next Obligation. now apply prf. Qed. Next Obligation. rewrite <-prf. discriminate. Qed. #[global] Program Instance prod_eq_dec `(A_dec : ∀ x y : A, Decision (x ≡ y)) `(B_dec : ∀ x y : B, Decision (x ≡ y)) : ∀ x y : A * B, Decision (x ≡ y) := λ x y, match A_dec (fst x) (fst y) with | left _ => match B_dec (snd x) (snd y) with left _ => left _ | right _ => right _ end | right _ => right _ end. Solve Obligations with (program_simpl; f_equal; firstorder). #[global] Program Instance and_dec `(P_dec : Decision P) `(Q_dec : Decision Q) : Decision (P ∧ Q) := match P_dec with | left _ => match Q_dec with left _ => left _ | right _ => right _ end | right _ => right _ end. Solve Obligations with (program_simpl; tauto). #[global] Program Instance or_dec `(P_dec : Decision P) `(Q_dec : Decision Q) : Decision (P ∨ Q) := match P_dec with | left _ => left _ | right _ => match Q_dec with left _ => left _ | right _ => right _ end end. Solve Obligations with (program_simpl; firstorder). #[global] Program Instance is_Some_dec `(x : option A) : Decision (is_Some x) := match x with | None => right _ | Some _ => left _ end. #[global] Program Instance is_None_dec `(x : option A) : Decision (is_None x) := match x with | None => left _ | Some _ => right _ end. #[global] Program Instance option_eq_dec `(A_dec : ∀ x y : A, Decision (x ≡ y)) : ∀ x y : option A, Decision (x ≡ y) := λ x y, match x with | Some r => match y with | Some s => match A_dec r s with left _ => left _ | right _ => right _ end | None => right _ end | None => match y with | Some s => right _ | None => left _ end end. #[global] Program Instance True_dec: Decision True := left _. #[global] Program Instance False_dec: Decision False := right _. math-classes-8.19.0/misc/propholds.v000066400000000000000000000015331460576051100173400ustar00rootroot00000000000000Require Import MathClasses.interfaces.canonical_names. (* The following class is nice to parametrize instances by additional properties, for example: [∀ z, PropHolds (z ≠ 0) → LeftCancellation op z] This tool is very powerful as we can combine it with instances as: [∀ x y, PropHolds (x ≠ 0) → PropHolds (y ≠ 0) → PropHolds (x * y ≠ 0)] Of course, one should be very careful not to make too many instances as that could easily lead to a blow-up of the search space (or worse, cycles). *) Class PropHolds (P : Prop) := prop_holds: P. #[global] Hint Extern 0 (PropHolds _) => assumption : typeclass_instances. #[global] Instance: Proper (iff ==> iff) PropHolds. Proof. now repeat intro. Qed. Ltac solve_propholds := match goal with | [ |- PropHolds (?P) ] => apply _ | [ |- ?P ] => change (PropHolds P); apply _ end. math-classes-8.19.0/misc/setoid_tactics.v000066400000000000000000000037601460576051100203330ustar00rootroot00000000000000Require Import Coq.Setoids.Setoid MathClasses.interfaces.canonical_names. (* When math-classes is used as part of another development setoid_replace often uses an incorrect equality due to low cost instances of DefaultRelation. We provide mc_setoid_replace to ensure that (=) is used. *) Tactic Notation "mc_setoid_replace" constr(x) "with" constr(y) := setoidreplace ((=) x y) idtac. Tactic Notation "mc_setoid_replace" constr(x) "with" constr(y) "at" int_or_var_list(o) := setoidreplaceat ((=) x y) idtac o. Tactic Notation "mc_setoid_replace" constr(x) "with" constr(y) "in" hyp(id) := setoidreplacein ((=) x y) id idtac. Tactic Notation "mc_setoid_replace" constr(x) "with" constr(y) "in" hyp(id) "at" int_or_var_list(o) := setoidreplaceinat ((=) x y) id idtac o. Tactic Notation "mc_setoid_replace" constr(x) "with" constr(y) "by" tactic3(t) := setoidreplace ((=) x y) ltac:(t). Tactic Notation "mc_setoid_replace" constr(x) "with" constr(y) "at" int_or_var_list(o) "by" tactic3(t) := setoidreplaceat ((=) x y) ltac:(t) o. Tactic Notation "mc_setoid_replace" constr(x) "with" constr(y) "in" hyp(id) "by" tactic3(t) := setoidreplacein ((=) x y) id ltac:(t). Tactic Notation "mc_setoid_replace" constr(x) "with" constr(y) "in" hyp(id) "at" int_or_var_list(o) "by" tactic3(t) := setoidreplaceinat ((=) x y) id ltac:(t) o. Ltac setoid_subst := repeat (match goal with | E : ?x = ?e |- _ => is_var x; lazymatch e with context [x] => fail | _ => rewrite ?E in *; clear x E end | E : ?e = ?x |- _ => is_var x; lazymatch e with context [x] => fail | _ => rewrite <-?E in *; clear x E end | E : ?x ≡ ?e |- _ => is_var x; lazymatch e with context [x] => fail | _ => rewrite ?E in *; clear x E end | E : ?e ≡ ?x |- _ => is_var x; lazymatch e with context [x] => fail | _ => rewrite <-?E in *; clear x E end end). Ltac setoid_discriminate := repeat intro; exfalso; match goal with | E : _ = _ |- _ => solve [inversion E] | E : _ ≡ _ |- _ => discriminate E end. math-classes-8.19.0/misc/stdlib_hints.v000066400000000000000000000023521460576051100200140ustar00rootroot00000000000000Require Import Coq.Setoids.Setoid Coq.Classes.Morphisms Coq.Classes.RelationClasses. Require Import Coq.Unicode.Utf8. #[global] Hint Unfold Proper respectful. #[global] Hint Unfold Reflexive Symmetric Transitive. #[global] Hint Constructors PreOrder. (* These tactics automatically solve symmetry and transitivity problems, when the hypothesis are in the context. They should be added as hints like Hint Extern 4 (?x = ?x) => reflexivity. Hint Extern 4 (?x = ?y) => auto_symm. Hint Extern 4 (?x = ?y) => auto_trans. once the appropriate head constants are defined. *) Ltac auto_symm := match goal with [ H: ?R ?x ?y |- ?R ?y ?x] => apply (symmetry H) end. Ltac auto_trans := match goal with [ H: ?R ?x ?y, I: ?R ?y ?z |- ?R ?x ?z] => apply (transitivity H I) end. (* These tactics make handling sig types slightly easier. *) Ltac exist_hyp := match goal with [ H : sig ?P |- ?P _ ] => exact (proj2_sig H) end. #[global] Hint Extern 0 => exist_hyp: core typeclass_instances. Ltac exist_proj1 := match goal with | [ |- context [proj1_sig ?x] ] => apply (proj2_sig x) end. #[global] Hint Extern 0 => exist_proj1: core typeclass_instances. math-classes-8.19.0/misc/util.v000066400000000000000000000055611460576051100163100ustar00rootroot00000000000000Require Import Coq.Program.Program Coq.Classes.Morphisms Coq.Setoids.Setoid MathClasses.interfaces.canonical_names. Section pointwise_dependent_relation. Context A (B: A → Type) (R: ∀ a, relation (B a)). Definition pointwise_dependent_relation: relation (∀ a, B a) := λ f f', ∀ a, R _ (f a) (f' a). Global Instance pdr_equiv `{∀ a, Equivalence (R a)}: Equivalence pointwise_dependent_relation. Proof. firstorder. Qed. End pointwise_dependent_relation. Definition iffT (A B: Type): Type := prod (A → B) (B → A). Class NonEmpty (A : Type) : Prop := non_empty : inhabited A. Class NonEmptyT (A : Type) : Type := non_emptyT : A. Definition uncurry {A B C} (f: A → B → C) (p: A * B): C := f (fst p) (snd p). Definition is_sole `{Equiv T} (P: T → Prop) (x: T) : Prop := P x ∧ ∀ y, P y → y = x. Definition DN (T: Type): Prop := (T → False) → False. Class Stable P := stable: DN P → P. (* TODO: include useful things from corn/logic/Stability.v and move to separate file *) Class Obvious (T : Type) := obvious: T. Section obvious. Context (A B C: Type). Global Instance: Obvious (A → A) := id. Global Instance: Obvious (False → A) := False_rect _. Global Instance: Obvious (A → A + B) := inl. Global Instance: Obvious (A → B + A) := inr. Global Instance obvious_sum_src `{Obvious (A → C)} `{Obvious (B → C)}: Obvious (A+B → C). Proof. repeat intro. intuition. Defined. Global Instance obvious_sum_dst_l `{Obvious (A → B)}: Obvious (A → B+C). Proof. repeat intro. intuition. Defined. Global Instance obvious_sum_dst_r `{Obvious (A → B)}: Obvious (A → C+B). Proof. repeat intro. intuition. Defined. End obvious. Lemma not_symmetry `{Symmetric A R} (x y: A): ¬R x y → ¬R y x. Proof. firstorder. Qed. (* Also see Coq bug #2358. A totally different approach would be to define negated relations such as inequality as separate relations rather than notations, so that the existing [symmetry] will work for them. However, this most likely breaks other things. *) Lemma biinduction_iff `{Biinduction R} (P1 : Prop) (P2 : R → Prop) (P2_proper : Proper ((=) ==> iff) P2) : (P1 ↔ P2 0) → (∀ n, P2 n ↔ P2 (1 + n)) → ∀ n, P1 ↔ P2 n. Proof. intros ? ?. apply biinduction; [solve_proper | easy | firstorder]. Qed. (* Isn't this in the stdlib? *) Definition is_Some `(x : option A) := match x with | None => False | Some _ => True end. Lemma is_Some_def `(x : option A) : is_Some x ↔ ∃ y, x ≡ Some y. Proof. unfold is_Some. destruct x; firstorder (eauto; discriminate). Qed. Definition is_None `(x : option A) := match x with | None => True | Some _ => False end. Lemma is_None_def `(x : option A) : is_None x ↔ x ≡ None. Proof. unfold is_None. destruct x; firstorder discriminate. Qed. Lemma None_ne_Some `(x : option A) y : x ≡ None → x ≢ Some y. Proof. congruence. Qed. math-classes-8.19.0/misc/workaround_tactics.v000066400000000000000000000006331460576051100212330ustar00rootroot00000000000000Global Ltac posed_rewrite t := let TEMP:=fresh in pose proof t as TEMP; simpl TEMP; rewrite TEMP; clear TEMP. (* Workaround around Coq bug, probably same as #2185. *) Global Ltac apply_simplified x := let TEMP:=fresh in generalize x; simpl; intro TEMP; apply TEMP; clear TEMP. Tactic Notation "posed_rewrite" "<-" constr(t) := let TEMP:= fresh in pose proof t as TEMP; simpl TEMP; rewrite <-TEMP; clear TEMP. math-classes-8.19.0/misc/workarounds.v000066400000000000000000000030511460576051100177010ustar00rootroot00000000000000Require Import MathClasses.interfaces.canonical_names. Require Import Coq.Classes.Equivalence Coq.Classes.Morphisms Coq.Classes.RelationClasses. (* Remove some duplicate / obsolete instances *) #[global] Remove Hints Equivalence_Reflexive equiv_reflexive Equivalence_Symmetric equiv_symmetric Equivalence_Transitive equiv_transitive : typeclass_instances. (* And re-insert the required ones with a low cost *) #[global] Hint Extern 0 (Reflexive _) => apply @Equivalence_Reflexive : typeclass_instances. #[global] Hint Extern 0 (Symmetric _) => apply @Equivalence_Symmetric : typeclass_instances. #[global] Hint Extern 0 (Transitive _) => apply @Equivalence_Transitive : typeclass_instances. (* (* We don't want Coq trying to prove e.g. transitivity of an arbitrary relation R by proving that R is a StrictOrder. Therefore we ensure that Coq only attempts so if R is actually an instance of Lt. *) Remove Hints StrictOrder_Transitive PreOrder_Reflexive PreOrder_Transitive PER_Symmetric PER_Transitive : typeclass_instances. Hint Extern 0 (Transitive (<)) => apply @StrictOrder_Transitive : typeclass_instances. Hint Extern 0 (Reflexive (≤)) => apply @PreOrder_Reflexive : typeclass_instances. Hint Extern 0 (Transitive (≤)) => apply @PreOrder_Transitive : typeclass_instances. *) (* It seems that Coq takes an insane number of steps to prove that an equivalence relation is Proper. This instance should decrease the number of performed steps. *) #[global] Instance equivalence_proper `{Equivalence A R} : Proper (R ==> R ==> iff) R | 0 := _. math-classes-8.19.0/orders/000077500000000000000000000000001460576051100155005ustar00rootroot00000000000000math-classes-8.19.0/orders/dec_fields.v000066400000000000000000000060441460576051100177540ustar00rootroot00000000000000Require Import Coq.Relations.Relation_Definitions Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.theory.rings MathClasses.theory.dec_fields. Require Export MathClasses.orders.rings. Section contents. Context `{DecField F} `{Apart F} `{!TrivialApart F} `{!FullPseudoSemiRingOrder Fle Flt} `{∀ x y : F, Decision (x = y)}. Add Ring F : (stdlib_ring_theory F). Instance pos_dec_recip_compat x : PropHolds (0 < x) → PropHolds (0 < /x). Proof. intros E. apply (strictly_order_reflecting (x *.)). rewrite dec_recip_inverse by now apply orders.lt_ne_flip. rewrite mult_0_r. solve_propholds. Qed. Instance nonneg_dec_recip_compat x : PropHolds (0 ≤ x) → PropHolds (0 ≤ /x). Proof. intros E. red. destruct (decide (x = 0)) as [E2 | E2]. now rewrite E2, dec_recip_0. apply lt_le. apply pos_dec_recip_compat. apply lt_iff_le_ne. split. easy. now apply not_symmetry. Qed. Lemma neg_dec_recip_compat x : x < 0 → /x < 0. Proof. intros. apply flip_neg_negate. rewrite dec_recip_negate. apply pos_dec_recip_compat. now apply flip_neg_negate. Qed. Lemma nonpos_dec_recip_compat x : x ≤ 0 → /x ≤ 0. Proof. intros. apply flip_nonpos_negate. rewrite dec_recip_negate. apply nonneg_dec_recip_compat. now apply flip_nonpos_negate. Qed. Lemma flip_le_dec_recip x y : 0 < y → y ≤ x → /x ≤ /y. Proof with trivial. intros E1 E2. apply (order_reflecting_pos (.*.) x)... now apply lt_le_trans with y. rewrite dec_recip_inverse. apply (order_reflecting_pos (.*.) y)... rewrite (commutativity x), associativity, dec_recip_inverse. now ring_simplify. now apply lt_ne_flip. apply lt_ne_flip. now apply lt_le_trans with y. Qed. Lemma flip_le_dec_recip_l x y : 0 < y → /y ≤ x → /x ≤ y. Proof with trivial. intros E1 E2. rewrite <-(dec_recip_involutive y). apply flip_le_dec_recip... now apply pos_dec_recip_compat. Qed. Lemma flip_le_dec_recip_r x y : 0 < y → y ≤ /x → x ≤ /y. Proof. intros E1 E2. rewrite <-(dec_recip_involutive x). now apply flip_le_dec_recip. Qed. Lemma flip_lt_dec_recip x y : 0 < y → y < x → /x < /y. Proof. intros E1 E2. assert (0 < x) by now transitivity y. apply (strictly_order_reflecting (x *.)). rewrite dec_recip_inverse. apply (strictly_order_reflecting (y *.)). rewrite (commutativity x), associativity, dec_recip_inverse. now ring_simplify. now apply lt_ne_flip. now apply lt_ne_flip. Qed. Lemma flip_lt_dec_recip_l x y : 0 < y → /y < x → /x < y. Proof. intros E1 E2. rewrite <-(dec_recip_involutive y). apply flip_lt_dec_recip; trivial. now apply pos_dec_recip_compat. Qed. Lemma flip_lt_dec_recip_r x y : 0 < y → y < /x → x < /y. Proof. intros E1 E2. rewrite <-(dec_recip_involutive x). now apply flip_lt_dec_recip. Qed. End contents. (* Due to bug #2528 *) #[global] Hint Extern 12 (PropHolds (0 ≤ _)) => eapply @nonneg_dec_recip_compat : typeclass_instances. #[global] Hint Extern 12 (PropHolds (0 < _)) => eapply @pos_dec_recip_compat : typeclass_instances. math-classes-8.19.0/orders/integers.v000066400000000000000000000106401460576051100175100ustar00rootroot00000000000000Require MathClasses.theory.integers MathClasses.theory.int_abs. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.integers MathClasses.interfaces.naturals MathClasses.interfaces.additional_operations MathClasses.interfaces.orders MathClasses.implementations.natpair_integers MathClasses.orders.rings. Require Export MathClasses.orders.nat_int. Section integers. Context `{Integers Z} `{Apart Z} `{!TrivialApart Z} `{!FullPseudoSemiRingOrder Zle Zlt}. Add Ring Z : (rings.stdlib_ring_theory Z). Add Ring nat : (rings.stdlib_semiring_theory nat). Lemma induction (P: Z → Prop) `{!Proper ((=) ==> iff) P}: P 0 → (∀ n, 0 ≤ n → P n → P (1 + n)) → (∀ n, n ≤ 0 → P n → P (n - 1)) → ∀ n, P n. Proof. intros P0 Psuc1 Psuc2 n. destruct (int_abs_sig Z nat n) as [[a A]|[a A]]. rewrite <-A. clear A. revert a. apply naturals.induction. solve_proper. now rewrite rings.preserves_0. intros m E. rewrite rings.preserves_plus, rings.preserves_1. apply Psuc1. apply to_semiring_nonneg. easy. rewrite <-(groups.negate_involutive n), <-A. clear A. revert a. apply naturals.induction. solve_proper. now rewrite rings.preserves_0, rings.negate_0. intros m E. rewrite rings.preserves_plus, rings.preserves_1. rewrite rings.negate_plus_distr, commutativity. apply Psuc2. apply naturals.negate_to_ring_nonpos. easy. Qed. Lemma induction_nonneg (P: Z → Prop) `{!Proper ((=) ==> iff) P}: P 0 → (∀ n, 0 ≤ n → P n → P (1 + n)) → ∀ n, 0 ≤ n → P n. Proof. intros P0 PS. refine (induction _ _ _ _); auto. solve_proper. intros n E1 ? E2. destruct (rings.is_ne_0 1). apply (antisymmetry (≤)). apply (order_reflecting ((n - 1) +)). ring_simplify. now transitivity 0. transitivity (n - 1). easy. apply (order_reflecting (1 +)). ring_simplify. transitivity 0. easy. apply semirings.le_0_2. Qed. Global Instance: Biinduction Z. Proof. intros P ? P0 Psuc. apply induction; trivial. firstorder. intros. apply Psuc. now setoid_replace (1 + (n - 1)) with n by ring. Qed. Global Program Instance: ∀ x y: Z, Decision (x ≤ y) | 10 := λ x y, match decide (integers_to_ring _ (SRpair nat) x ≤ integers_to_ring _ (SRpair nat) y) with | left E => left _ | right E => right _ end. Next Obligation. now apply (order_reflecting (integers_to_ring _ (SRpair nat))). Qed. End integers. (* A default order on the integers *) #[global] Instance int_le `{Integers Z} : Le Z | 10 := λ x y, ∃ z, y = x + naturals_to_semiring nat Z z. #[global] Instance int_lt `{Integers Z} : Lt Z | 10 := dec_lt. Section default_order. Context `{Integers Int} `{Apart Int} `{!TrivialApart Int}. Add Ring Int2 : (rings.stdlib_ring_theory Int). Instance: Proper ((=) ==> (=) ==> iff) int_le. Proof. intros x1 y1 E1 x2 y2 E2. split; intros [z p]; exists z. now rewrite <-E1, <-E2. now rewrite E1, E2. Qed. Instance: PartialOrder int_le. Proof. repeat (split; try apply _). intros x. exists 0. rewrite rings.preserves_0. ring. intros x y z [a A] [b B]. exists (a + b). now rewrite rings.preserves_plus, associativity, <-A, B. intros x y [a A] [b B]. destruct (naturals.zero_sum a b) as [E1 E2]. apply (injective (naturals_to_semiring nat Int)). rewrite rings.preserves_plus, rings.preserves_0. apply (left_cancellation (+) x). rewrite B at 2. rewrite A. ring. rewrite A, B, E1, E2, rings.preserves_0. ring. Qed. Instance: SemiRingOrder int_le. Proof. apply from_ring_order. repeat (split; try apply _). intros x y [a A]. exists a. rewrite A. ring. intros x y [a A] [b B]. exists (a * b). rewrite A, B, rings.preserves_mult. ring. Qed. Notation i_to_r := (integers_to_ring Int (SRpair nat)). Instance: TotalRelation int_le. Proof. assert (∀ x y, i_to_r x ≤ i_to_r y → x ≤ y) as P. intros x y E. destruct (decompose_le E) as [a [A B]]. exists (pos a ∸ neg a). apply (injective i_to_r). rewrite rings.preserves_plus, B. clear B. apply sm_proper. rewrite (naturals.to_semiring_twice _ _ SRpair_inject). unfold equiv, SRpair_equiv, le, SRpair_le in *. simpl in *. rewrite right_identity, cut_minus_le. reflexivity. now rewrite rings.plus_0_l, rings.plus_0_r in A. intros x y. now destruct (total (≤) (i_to_r x) (i_to_r y)); [left|right]; eapply P. Qed. Global Instance: FullPseudoSemiRingOrder int_le int_lt. Proof. now apply dec_full_pseudo_srorder. Qed. End default_order. math-classes-8.19.0/orders/lattices.v000066400000000000000000000300641460576051100175020ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.orders.maps MathClasses.theory.lattices. (* We prove that the algebraic definition of a lattice corresponds to the order theoretic one. Note that we do not make any of these instances global, because that would cause loops. *) Section join_semilattice_order. Context `{JoinSemiLatticeOrder L}. Instance: Setoid L := po_setoid. Lemma join_ub_3_r x y z : z ≤ x ⊔ y ⊔ z. Proof. now apply join_ub_r. Qed. Lemma join_ub_3_m x y z : y ≤ x ⊔ y ⊔ z. Proof. transitivity (x ⊔ y). now apply join_ub_r. now apply join_ub_l. Qed. Lemma join_ub_3_l x y z : x ≤ x ⊔ y ⊔ z. Proof. transitivity (x ⊔ y); now apply join_ub_l. Qed. Lemma join_ub_3_assoc_l x y z : x ≤ x ⊔ (y ⊔ z). Proof. now apply join_ub_l. Qed. Lemma join_ub_3_assoc_m x y z : y ≤ x ⊔ (y ⊔ z). Proof. transitivity (y ⊔ z). now apply join_ub_l. now apply join_ub_r. Qed. Lemma join_ub_3_assoc_r x y z : z ≤ x ⊔ (y ⊔ z). Proof. transitivity (y ⊔ z); now apply join_ub_r. Qed. Instance: Proper ((=) ==> (=) ==> (=)) (⊔). Proof. intros ? ? E1 ? ? E2. apply (antisymmetry (≤)); apply join_lub. rewrite E1. now apply join_ub_l. rewrite E2. now apply join_ub_r. rewrite <-E1. now apply join_ub_l. rewrite <-E2. now apply join_ub_r. Qed. Instance join_sl_order_join_sl: JoinSemiLattice L. Proof. repeat (split; try apply _). intros x y z. apply (antisymmetry (≤)). apply join_lub. now apply join_ub_3_l. apply join_lub. now apply join_ub_3_m. now apply join_ub_3_r. apply join_lub. apply join_lub. now apply join_ub_3_assoc_l. now apply join_ub_3_assoc_m. now apply join_ub_3_assoc_r. intros x y. apply (antisymmetry (≤)); apply join_lub; first [apply join_ub_l | try apply join_ub_r]. intros x. red. apply (antisymmetry (≤)). now apply join_lub. now apply join_ub_l. Qed. Lemma join_le_compat_r x y z : z ≤ x → z ≤ x ⊔ y. Proof. intros E. transitivity x. easy. apply join_ub_l. Qed. Lemma join_le_compat_l x y z : z ≤ y → z ≤ x ⊔ y. Proof. intros E. rewrite commutativity. now apply join_le_compat_r. Qed. Lemma join_l x y : y ≤ x → x ⊔ y = x. Proof. intros E. apply (antisymmetry (≤)). now apply join_lub. apply join_ub_l. Qed. Lemma join_r x y : x ≤ y → x ⊔ y = y. Proof. intros E. rewrite commutativity. now apply join_l. Qed. Lemma join_sl_le_spec x y : x ≤ y ↔ x ⊔ y = y. Proof. split; intros E. now apply join_r. rewrite <-E. now apply join_ub_l. Qed. Global Instance: ∀ z, OrderPreserving (z ⊔). Proof. intros. repeat (split; try apply _). intros. apply join_lub. now apply join_ub_l. now apply join_le_compat_l. Qed. Global Instance: ∀ z, OrderPreserving (⊔ z). Proof. intros. apply maps.order_preserving_flip. Qed. Lemma join_le_compat x₁ x₂ y₁ y₂ : x₁ ≤ x₂ → y₁ ≤ y₂ → x₁ ⊔ y₁ ≤ x₂ ⊔ y₂. Proof. intros E1 E2. transitivity (x₁ ⊔ y₂). now apply (order_preserving (x₁ ⊔)). now apply (order_preserving (⊔ y₂)). Qed. Lemma join_le x y z : x ≤ z → y ≤ z → x ⊔ y ≤ z. Proof. intros. rewrite <-(idempotency (⊔) z). now apply join_le_compat. Qed. End join_semilattice_order. Section bounded_join_semilattice. Context `{JoinSemiLatticeOrder L} `{Bottom L} `{!BoundedJoinSemiLattice L}. Lemma above_bottom x : ⊥ ≤ x. Proof. rewrite join_sl_le_spec. now rewrite left_identity. Qed. Lemma below_bottom x : x ≤ ⊥ → x = ⊥. Proof. rewrite join_sl_le_spec. now rewrite right_identity. Qed. End bounded_join_semilattice. Section meet_semilattice_order. Context `{MeetSemiLatticeOrder L}. Instance: Setoid L := po_setoid. Lemma meet_lb_3_r x y z : x ⊓ y ⊓ z ≤ z. Proof. now apply meet_lb_r. Qed. Lemma meet_lb_3_m x y z : x ⊓ y ⊓ z ≤ y. Proof. transitivity (x ⊓ y). now apply meet_lb_l. now apply meet_lb_r. Qed. Lemma meet_lb_3_l x y z : x ⊓ y ⊓ z ≤ x. Proof. transitivity (x ⊓ y); now apply meet_lb_l. Qed. Lemma meet_lb_3_assoc_l x y z : x ⊓ (y ⊓ z) ≤ x. Proof. now apply meet_lb_l. Qed. Lemma meet_lb_3_assoc_m x y z : x ⊓ (y ⊓ z) ≤ y. Proof. transitivity (y ⊓ z). now apply meet_lb_r. now apply meet_lb_l. Qed. Lemma meet_lb_3_assoc_r x y z : x ⊓ (y ⊓ z) ≤ z. Proof. transitivity (y ⊓ z); now apply meet_lb_r. Qed. Instance: Proper ((=) ==> (=) ==> (=)) (⊓). Proof. intros ? ? E1 ? ? E2. apply (antisymmetry (≤)); apply meet_glb. rewrite <-E1. now apply meet_lb_l. rewrite <-E2. now apply meet_lb_r. rewrite E1. now apply meet_lb_l. rewrite E2. now apply meet_lb_r. Qed. Instance meet_sl_order_meet_sl: MeetSemiLattice L. Proof. repeat (split; try apply _). intros x y z. apply (antisymmetry (≤)). apply meet_glb. apply meet_glb. now apply meet_lb_3_assoc_l. now apply meet_lb_3_assoc_m. now apply meet_lb_3_assoc_r. apply meet_glb. now apply meet_lb_3_l. apply meet_glb. now apply meet_lb_3_m. now apply meet_lb_3_r. intros x y. apply (antisymmetry (≤)); apply meet_glb; first [apply meet_lb_l | try apply meet_lb_r]. intros x. red. apply (antisymmetry (≤)). now apply meet_lb_l. now apply meet_glb. Qed. Lemma meet_le_compat_r x y z : x ≤ z → x ⊓ y ≤ z. Proof. intros E. transitivity x. apply meet_lb_l. easy. Qed. Lemma meet_le_compat_l x y z : y ≤ z → x ⊓ y ≤ z. Proof. intros E. rewrite commutativity. now apply meet_le_compat_r. Qed. Lemma meet_l x y : x ≤ y → x ⊓ y = x. Proof. intros E. apply (antisymmetry (≤)). apply meet_lb_l. now apply meet_glb. Qed. Lemma meet_r x y : y ≤ x → x ⊓ y = y. Proof. intros E. rewrite commutativity. now apply meet_l. Qed. Lemma meet_sl_le_spec x y : x ≤ y ↔ x ⊓ y = x. Proof. split; intros E. now apply meet_l. rewrite <-E. now apply meet_lb_r. Qed. Global Instance: ∀ z, OrderPreserving (z ⊓). Proof. intros. repeat (split; try apply _). intros. apply meet_glb. now apply meet_lb_l. now apply meet_le_compat_l. Qed. Global Instance: ∀ z, OrderPreserving (⊓ z). Proof. intros. apply maps.order_preserving_flip. Qed. Lemma meet_le_compat x₁ x₂ y₁ y₂ : x₁ ≤ x₂ → y₁ ≤ y₂ → x₁ ⊓ y₁ ≤ x₂ ⊓ y₂. Proof. intros E1 E2. transitivity (x₁ ⊓ y₂). now apply (order_preserving (x₁ ⊓)). now apply (order_preserving (⊓ y₂)). Qed. Lemma meet_le x y z : z ≤ x → z ≤ y → z ≤ x ⊓ y. Proof. intros. rewrite <-(idempotency (⊓) z). now apply meet_le_compat. Qed. End meet_semilattice_order. Section lattice_order. Context `{LatticeOrder L}. Instance: JoinSemiLattice L := join_sl_order_join_sl. Instance: MeetSemiLattice L := meet_sl_order_meet_sl. Instance: Absorption (⊓) (⊔). Proof. intros x y. apply (antisymmetry (≤)). now apply meet_lb_l. apply meet_le. easy. now apply join_ub_l. Qed. Instance: Absorption (⊔) (⊓). Proof. intros x y. apply (antisymmetry (≤)). apply join_le. easy. now apply meet_lb_l. now apply join_ub_l. Qed. Instance lattice_order_lattice: Lattice L. Proof. split; try apply _. Qed. Lemma meet_join_distr_l_le x y z : (x ⊓ y) ⊔ (x ⊓ z) ≤ x ⊓ (y ⊔ z). Proof. apply meet_le. apply join_le; now apply meet_lb_l. apply join_le. transitivity y. apply meet_lb_r. apply join_ub_l. transitivity z. apply meet_lb_r. apply join_ub_r. Qed. Lemma join_meet_distr_l_le x y z : x ⊔ (y ⊓ z) ≤ (x ⊔ y) ⊓ (x ⊔ z). Proof. apply meet_le. apply join_le. now apply join_ub_l. transitivity y. apply meet_lb_l. apply join_ub_r. apply join_le. apply join_ub_l. transitivity z. apply meet_lb_r. apply join_ub_r. Qed. End lattice_order. Definition default_join_sl_le `{JoinSemiLattice L} : Le L := λ x y, x ⊔ y = y. Section join_sl_order_alt. Context `{JoinSemiLattice L} `{Le L} (le_correct : ∀ x y, x ≤ y ↔ x ⊔ y = y). Lemma alt_Build_JoinSemiLatticeOrder : JoinSemiLatticeOrder (≤). Proof. split; try (split; try apply _). intros ?? E1 ?? E2. now rewrite !le_correct, E1, E2. split. intros ?. rewrite !le_correct. now apply (idempotency _ _). intros ? ? ?. rewrite !le_correct. intros E1 E2. now rewrite <-E2, associativity, E1. intros ? ?. rewrite !le_correct. intros E1 E2. now rewrite <-E1, commutativity, <-E2 at 1. intros ? ?. now rewrite le_correct, associativity, (idempotency _ _). intros ? ?. now rewrite le_correct, commutativity, <-associativity, (idempotency _ _). intros ? ? ?. rewrite !le_correct. intros E1 E2. now rewrite <-associativity, E2. Qed. End join_sl_order_alt. Definition default_meet_sl_le `{MeetSemiLattice L} : Le L := λ x y, x ⊓ y = x. Section meet_sl_order_alt. Context `{MeetSemiLattice L} `{Le L} (le_correct : ∀ x y, x ≤ y ↔ x ⊓ y = x). Lemma alt_Build_MeetSemiLatticeOrder : MeetSemiLatticeOrder (≤). Proof. split; try (split; try apply _). intros ?? E1 ?? E2. now rewrite !le_correct, E1, E2. split. intros ?. rewrite !le_correct. now apply (idempotency _ _). intros ? ? ?. rewrite !le_correct. intros E1 E2. now rewrite <-E1, <-associativity, E2. intros ? ?. rewrite !le_correct. intros E1 E2. now rewrite <-E2, commutativity, <-E1 at 1. intros ? ?. now rewrite le_correct, commutativity, associativity, (idempotency _ _). intros ? ?. now rewrite le_correct, <-associativity, (idempotency _ _). intros ? ? ?. rewrite !le_correct. intros E1 E2. now rewrite associativity, E1. Qed. End meet_sl_order_alt. Section join_order_preserving. Context `{JoinSemiLatticeOrder L} `{JoinSemiLatticeOrder K} (f : L → K) `{!JoinSemiLattice_Morphism f}. Local Existing Instance join_sl_order_join_sl. Lemma join_sl_mor_preserving: OrderPreserving f. Proof. repeat (split; try apply _). intros ??. rewrite 2!join_sl_le_spec, <-preserves_join. intros E. now rewrite E. Qed. Lemma join_sl_mor_reflecting `{!Injective f}: OrderReflecting f. Proof. repeat (split; try apply _). intros ??. rewrite 2!join_sl_le_spec, <-preserves_join. intros. now apply (injective f). Qed. End join_order_preserving. Section meet_order_preserving. Context `{MeetSemiLatticeOrder L} `{MeetSemiLatticeOrder K} (f : L → K) `{!MeetSemiLattice_Morphism f}. Local Existing Instance meet_sl_order_meet_sl. Lemma meet_sl_mor_preserving: OrderPreserving f. Proof. repeat (split; try apply _). intros ??. rewrite 2!meet_sl_le_spec, <-preserves_meet. intros E. now rewrite E. Qed. Lemma meet_sl_mor_reflecting `{!Injective f}: OrderReflecting f. Proof. repeat (split; try apply _). intros ??. rewrite 2!meet_sl_le_spec, <-preserves_meet. intros. now apply (injective f). Qed. End meet_order_preserving. Section order_preserving_join_sl_mor. Context `{JoinSemiLatticeOrder L} `{JoinSemiLatticeOrder K} `{!TotalOrder (_ : Le L)} `{!TotalOrder (_ : Le K)} `{!OrderPreserving (f : L → K)}. Local Existing Instance join_sl_order_join_sl. Local Existing Instance order_morphism_mor. Lemma order_preserving_join_sl_mor: JoinSemiLattice_Morphism f. Proof. repeat (split; try apply _). intros x y. case (total (≤) x y); intros E. rewrite 2!join_r; try easy. now apply (order_preserving _). rewrite 2!join_l; try easy. now apply (order_preserving _). Qed. End order_preserving_join_sl_mor. Section order_preserving_meet_sl_mor. Context `{MeetSemiLatticeOrder L} `{MeetSemiLatticeOrder K} `{!TotalOrder (_ : Le L)} `{!TotalOrder (_ : Le K)} `{!OrderPreserving (f : L → K)}. Local Existing Instance meet_sl_order_meet_sl. Local Existing Instance order_morphism_mor. Lemma order_preserving_meet_sl_mor: SemiGroup_Morphism f. Proof. repeat (split; try apply _). intros x y. case (total (≤) x y); intros E. rewrite 2!meet_l; try easy. now apply (order_preserving _). rewrite 2!meet_r; try easy. now apply (order_preserving _). Qed. End order_preserving_meet_sl_mor. math-classes-8.19.0/orders/maps.v000066400000000000000000000315511460576051100166340ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.orders.orders MathClasses.theory.setoids MathClasses.theory.strong_setoids. Local Existing Instance order_morphism_po_a. Local Existing Instance order_morphism_po_b. Local Existing Instance strict_order_morphism_so_a. Local Existing Instance strict_order_morphism_so_b. Local Existing Instance order_morphism_mor. Local Existing Instance strict_order_morphism_mor. (* If a function between strict partial orders is order preserving (back), we can derive that it is strictly order preserving (back) *) Section strictly_order_preserving. Context `{FullPartialOrder A} `{FullPartialOrder B}. Global Instance strictly_order_preserving_inj `{!OrderPreserving (f : A → B)} `{!StrongInjective f} : StrictlyOrderPreserving f | 20. Proof. repeat (split; try apply _). intros x y. rewrite !lt_iff_le_apart. intros [E1 E2]. split. now apply (order_preserving f). now apply (strong_injective f). Qed. Global Instance strictly_order_reflecting_mor `{!OrderReflecting (f : A → B)} `{!StrongSetoid_Morphism f} : StrictlyOrderReflecting f | 20. Proof. repeat (split; try apply _). intros x y. rewrite !lt_iff_le_apart. intros [E1 E2]. split. now apply (order_reflecting f). now apply (strong_extensionality f). Qed. End strictly_order_preserving. (* For structures with a trivial apartness relation we have a stronger result of the above *) Section strictly_order_preserving_dec. Context `{FullPartialOrder A} `{!TrivialApart A} `{FullPartialOrder B} `{!TrivialApart B}. Local Existing Instance strict_po_setoid. Global Instance dec_strictly_order_preserving_inj `{!OrderPreserving (f : A → B)} `{!Injective f} : StrictlyOrderPreserving f | 19. Proof. pose proof (dec_strong_injective f). apply _. Qed. Global Instance dec_strictly_order_reflecting_mor `{!OrderReflecting (f : A → B)} : StrictlyOrderReflecting f | 19. Proof. pose proof (order_morphism_mor f). pose proof (dec_strong_morphism f). apply _. Qed. End strictly_order_preserving_dec. Section pseudo_injective. Context `{PseudoOrder A} `{PseudoOrder B}. Local Existing Instance pseudo_order_setoid. Instance pseudo_order_embedding_ext `{!StrictOrderEmbedding (f : A → B)} : StrongSetoid_Morphism f. Proof. split; try apply _. intros x y. rewrite !apart_iff_total_lt. intros [|]; [left | right]; now apply (strictly_order_reflecting f). Qed. Lemma pseudo_order_embedding_inj `{!StrictOrderEmbedding (f : A → B)} : StrongInjective f. Proof. split; try apply _. intros x y. rewrite !apart_iff_total_lt. intros [|]; [left | right]; now apply (strictly_order_preserving f). Qed. End pseudo_injective. (* If a function between pseudo partial orders is strictly order preserving (back), we can derive that it is order preserving (back) *) Section full_pseudo_strictly_preserving. Context `{FullPseudoOrder A} `{FullPseudoOrder B}. Local Existing Instance pseudo_order_setoid. Lemma full_pseudo_order_preserving `{!StrictlyOrderReflecting (f : A → B)} : OrderPreserving f. Proof. pose proof (strict_order_morphism_mor f). repeat (split; try apply _). intros x y. rewrite !le_iff_not_lt_flip. intros E1 E2. apply E1. now apply (strictly_order_reflecting f). Qed. Lemma full_pseudo_order_reflecting `{!StrictlyOrderPreserving (f : A → B)} : OrderReflecting f. Proof. pose proof (strict_order_morphism_mor f). repeat (split; try apply _). intros x y. rewrite !le_iff_not_lt_flip. intros E1 E2. apply E1. now apply (strictly_order_preserving f). Qed. End full_pseudo_strictly_preserving. (* Some helper lemmas to easily transform order preserving instances. *) Section order_preserving_ops. Context `{Equiv R} `{Le R} `{Lt R}. Lemma order_preserving_flip `{!Commutative op} `{!Proper ((=) ==> (=) ==> (=)) op} `{!OrderPreserving (op z)} : OrderPreserving (λ y, op y z). Proof. pose proof (order_morphism_mor (op z)). pose proof (setoidmor_a (op z)). repeat (split; try apply _). solve_proper. intros x y E. rewrite 2!(commutativity _ z). now apply order_preserving. Qed. Lemma strictly_order_preserving_flip `{!Commutative op} `{!Proper ((=) ==> (=) ==> (=)) op} `{!StrictlyOrderPreserving (op z)} : StrictlyOrderPreserving (λ y, op y z). Proof. pose proof (strict_order_morphism_mor (op z)). pose proof (setoidmor_a (op z)). repeat (split; try apply _). solve_proper. intros x y E. rewrite 2!(commutativity _ z). now apply strictly_order_preserving. Qed. Lemma order_reflecting_flip `{!Commutative op} `{!Proper ((=) ==> (=) ==> (=)) op} `{!OrderReflecting (op z) } : OrderReflecting (λ y, op y z). Proof. pose proof (order_morphism_mor (op z)). pose proof (setoidmor_a (op z)). repeat (split; try apply _). solve_proper. intros x y E. apply (order_reflecting (op z)). now rewrite 2!(commutativity z). Qed. Lemma strictly_order_reflecting_flip `{!Commutative op} `{!Proper ((=) ==> (=) ==> (=)) op} `{!StrictlyOrderReflecting (op z) } : StrictlyOrderReflecting (λ y, op y z). Proof. pose proof (strict_order_morphism_mor (op z)). pose proof (setoidmor_a (op z)). repeat (split; try apply _). solve_proper. intros x y E. apply (strictly_order_reflecting (op z)). now rewrite 2!(commutativity z). Qed. Lemma order_preserving_nonneg (op : R → R → R) `{!Zero R} `{∀ z, PropHolds (0 ≤ z) → OrderPreserving (op z)} z : 0 ≤ z → ∀ x y, x ≤ y → op z x ≤ op z y. Proof. firstorder. Qed. Lemma order_preserving_flip_nonneg (op : R → R → R) `{!Zero R} `{∀ z, PropHolds (0 ≤ z) → OrderPreserving (λ y, op y z)} z : 0 ≤ z → ∀ x y, x ≤ y → op x z ≤ op y z. Proof. firstorder. Qed. Lemma strictly_order_preserving_pos (op : R → R → R) `{!Zero R} `{∀ z, PropHolds (0 < z) → StrictlyOrderPreserving (op z)} z : 0 < z → ∀ x y, x < y → op z x < op z y. Proof. firstorder. Qed. Lemma strictly_order_preserving_flip_pos (op : R → R → R) `{!Zero R} `{∀ z, PropHolds (0 < z) → StrictlyOrderPreserving (λ y, op y z)} z : 0 < z → ∀ x y, x < y → op x z < op y z. Proof. firstorder. Qed. Lemma order_reflecting_pos (op : R → R → R) `{!Zero R} `{∀ z, PropHolds (0 < z) → OrderReflecting (op z)} z : 0 < z → ∀ x y, op z x ≤ op z y → x ≤ y. Proof. firstorder. Qed. Lemma order_reflecting_flip_pos (op : R → R → R) `{!Zero R} `{∀ z, PropHolds (0 < z) → OrderReflecting (λ y, op y z)} z : 0 < z → ∀ x y, op x z ≤ op y z → x ≤ y. Proof. firstorder. Qed. End order_preserving_ops. Lemma projected_partial_order `{Equiv A} `{Ale : Le A} `{Equiv B} `{Ble : Le B} (f : A → B) `{!Injective f} `{!PartialOrder Ble} : (∀ x y, x ≤ y ↔ f x ≤ f y) → PartialOrder Ale. Proof. pose proof (injective_mor f). pose proof (setoidmor_a f). pose proof (setoidmor_b f). intros P. split; try apply _. intros ? ? E1 ? ? E2. now rewrite 2!P, E1, E2. split. intros x. now apply P. intros x y z E1 E2. apply P. transitivity (f y); now apply P. intros x y E1 E2. apply (injective f). apply (antisymmetry (≤)); now apply P. Qed. Lemma projected_total_order `{Equiv A} `{Ale : Le A} `{Equiv B} `{Ble : Le B} (f : A → B) `{!TotalRelation Ble} : (∀ x y, x ≤ y ↔ f x ≤ f y) → TotalRelation Ale. Proof. intros P x y. destruct (total (≤) (f x) (f y)); [left | right]; now apply P. Qed. Lemma projected_strict_order `{Equiv A} `{Alt : Lt A} `{Equiv B} `{Blt : Lt B} (f : A → B) `{!StrictOrder Blt} : (∀ x y, x < y ↔ f x < f y) → StrictOrder Alt. Proof. intros P. split. intros x E. destruct (irreflexivity (<) (f x)). now apply P. intros x y z E1 E2. apply P. transitivity (f y); now apply P. Qed. Lemma projected_strict_setoid_order `{Equiv A} `{Alt : Lt A} `{Equiv B} `{Blt : Lt B} (f : A → B) `{!Setoid_Morphism f} `{!StrictSetoidOrder Blt} : (∀ x y, x < y ↔ f x < f y) → StrictSetoidOrder Alt. Proof. pose proof (setoidmor_a f). intros P. split; try apply _. intros ? ? E1 ? ? E2. now rewrite 2!P, E1, E2. now apply (projected_strict_order f). Qed. Lemma projected_pseudo_order `{Equiv A} `{Apart A} `{Alt : Lt A} `{Equiv B} `{Apart B} `{Blt : Lt B} (f : A → B) `{!StrongInjective f} `{!PseudoOrder Blt} : (∀ x y, x < y ↔ f x < f y) → PseudoOrder Alt. Proof. pose proof (strong_injective_mor f). pose proof (strong_setoidmor_a f). pose proof (strong_setoidmor_b f). intros P. split; try apply _. intros x y E. destruct (pseudo_order_antisym (f x) (f y)). split; now apply P. intros x y E z. apply P in E. destruct (cotransitive E (f z)); [left | right]; now apply P. intros x y; split; intros E. apply (strong_injective f) in E. apply apart_iff_total_lt in E. destruct E; [left | right]; now apply P. apply (strong_extensionality f). apply apart_iff_total_lt. destruct E; [left | right]; now apply P. Qed. Lemma projected_full_pseudo_order `{Equiv A} `{Apart A} `{Ale : Le A} `{Alt : Lt A} `{Equiv B} `{Apart B} `{Ble : Le B} `{Blt : Lt B} (f : A → B) `{!StrongInjective f} `{!FullPseudoOrder Ble Blt} : (∀ x y, x ≤ y ↔ f x ≤ f y) → (∀ x y, x < y ↔ f x < f y) → FullPseudoOrder Ale Alt. Proof. intros P1 P2. split. now apply (projected_pseudo_order f). intros x y; split; intros E. intros F. destruct (le_not_lt_flip (f y) (f x)); firstorder. apply P1. apply not_lt_le_flip. intros F. destruct E. now apply P2. Qed. #[global] Instance id_order_morphism `{PartialOrder A} : Order_Morphism (@id A). Proof. pose proof po_setoid. repeat (split; try apply _). Qed. #[global] Instance id_order_preserving `{PartialOrder A} : OrderPreserving (@id A). Proof. split; try apply _. easy. Qed. #[global] Instance id_order_reflecting `{PartialOrder A} : OrderReflecting (@id A). Proof. split; try apply _. easy. Qed. Section composition. Context `{Equiv A} `{Equiv B} `{Equiv C} `{Le A} `{Le B} `{Le C} (f : A → B) (g : B → C). Instance compose_order_morphism: Order_Morphism f → Order_Morphism g → Order_Morphism (g ∘ f). Proof. split; [ apply _ | apply (order_morphism_po_a f) | apply (order_morphism_po_b g) ]. Qed. Instance compose_order_preserving: OrderPreserving f → OrderPreserving g → OrderPreserving (g ∘ f). Proof. repeat (split; try apply _). intros. unfold compose. now do 2 apply (order_preserving _). Qed. Instance compose_order_reflecting: OrderReflecting f → OrderReflecting g → OrderReflecting (g ∘ f). Proof. split; try apply _. intros x y E. unfold compose in E. now do 2 apply (order_reflecting _) in E. Qed. Instance compose_order_embedding: OrderEmbedding f → OrderEmbedding g → OrderEmbedding (g ∘ f) := {}. End composition. #[global] Hint Extern 4 (Order_Morphism (_ ∘ _)) => class_apply @compose_order_morphism : typeclass_instances. #[global] Hint Extern 4 (OrderPreserving (_ ∘ _)) => class_apply @compose_order_preserving : typeclass_instances. #[global] Hint Extern 4 (OrderReflecting (_ ∘ _)) => class_apply @compose_order_reflecting : typeclass_instances. #[global] Hint Extern 4 (OrderEmbedding (_ ∘ _)) => class_apply @compose_order_embedding : typeclass_instances. Section propers. Context `{Equiv A} `{Equiv B} `{Le A} `{Le B}. Global Instance order_morphism_proper: Proper ((=) ==> iff) (@Order_Morphism A B _ _ _ _). Proof. assert (∀ (f g : A → B), g = f → Order_Morphism f → Order_Morphism g) as P. intros f g E [[? ? ?] ?]. split; auto. apply morphism_proper with f. easy. split; easy. firstorder. Qed. Global Instance order_preserving_proper: Proper ((=) ==> iff) (@OrderPreserving A B _ _ _ _). Proof. assert (∀ (f g : A → B), g = f → OrderPreserving f → OrderPreserving g) as P. intros f g E [[[? ?] ? ?] ?]. split. eapply order_morphism_proper; eauto. now repeat (split; try apply _). intros x y ?. rewrite (E x x), (E y y); now auto. firstorder. Qed. Global Instance order_reflecting_proper: Proper ((=) ==> iff) (@OrderReflecting A B _ _ _ _). Proof. assert (∀ (f g : A → B), g = f → OrderReflecting f → OrderReflecting g) as P. intros f g E [[[? ?] ? ?] ?]. split. eapply order_morphism_proper; eauto. now repeat (split; try apply _). intros x y F. rewrite (E x x), (E y y) in F; now auto. firstorder. Qed. Global Instance order_embedding_proper: Proper ((=) ==> iff) (@OrderEmbedding A B _ _ _ _). Proof. assert (∀ (f g : A → B), g = f → OrderEmbedding f → OrderEmbedding g) as P. intros f g E. split. eapply order_preserving_proper; eauto. now apply _. eapply order_reflecting_proper; eauto. now apply _. intros f g ?; split; intro E. apply P with f. destruct E as [[[[? ?]]]]. now symmetry. easy. now apply P with g. Qed. End propers. math-classes-8.19.0/orders/minmax.v000066400000000000000000000025641460576051100171670ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.orders.orders MathClasses.orders.lattices MathClasses.theory.setoids. Section contents. Context `{TotalOrder A} `{∀ x y: A, Decision (x ≤ y)}. Instance: Setoid A := po_setoid. Definition sort (x y: A) : A * A := if decide_rel (≤) x y then (x, y) else (y, x). Global Instance: Proper ((=) ==> (=) ==> (=)) sort. Proof. intros ? ? E1 ? ? E2. unfold sort. do 2 case (decide_rel _); simpl. firstorder. intros F ?. destruct F. now rewrite <-E1, <-E2. intros ? F. destruct F. now rewrite E1, E2. firstorder. Qed. Global Instance min: Meet A := λ x y, fst (sort x y). Global Instance max: Join A := λ x y, snd (sort x y). Global Instance: LatticeOrder (≤). Proof. repeat (split; try apply _); unfold join, meet, max, min, sort; intros; case (decide_rel _); try easy; now apply le_flip. Qed. Instance: LeftDistribute max min. Proof. intros x y z. unfold min, max, sort. repeat case (decide_rel _); simpl; try solve [intuition]. intros. apply (antisymmetry (≤)); [|easy]. now transitivity y; apply le_flip. intros. now apply (antisymmetry (≤)). Qed. Instance: Lattice A := lattice_order_lattice. Global Instance: DistributiveLattice A. Proof. repeat (split; try apply _). Qed. End contents. math-classes-8.19.0/orders/nat_int.v000066400000000000000000000117671460576051100173370ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals MathClasses.interfaces.orders MathClasses.theory.naturals MathClasses.theory.rings MathClasses.implementations.peano_naturals. Require Export MathClasses.orders.semirings. (* We axiomatize the order on the naturals and the integers as a non trivial pseudo semiring order that satisfies the biinduction principle. We prove some results that hold for the order on the naturals and the integers. In particular, we show that given another non trivial pseudo semiring order (that not necessarily has to satisfy the biinduction principle, for example the rationals or the reals), any morphism to it is an order embedding. *) Lemma to_semiring_nonneg `{FullPseudoSemiRingOrder N} `{!NaturalsToSemiRing N} `{!Naturals N} `{FullPseudoSemiRingOrder R} `{!SemiRing_Morphism (f : N → R)} n : 0 ≤ f n. Proof. pose proof (pseudo_srorder_semiring (A:=R)). revert n. apply naturals.induction. solve_proper. intros. now rewrite preserves_0. intros n E. rewrite preserves_plus, preserves_1. apply nonneg_plus_compat. solve_propholds. easy. Qed. Section nat_int_order. Context `{FullPseudoSemiRingOrder R} `{!Biinduction R} `{PropHolds (1 ≶ 0)}. Local Existing Instance pseudo_srorder_semiring. Add Ring R : (stdlib_semiring_theory R). (* Hmm, can we avoid using nat here? *) Lemma nat_int_to_semiring x : ∃ z, x = naturals_to_semiring nat R z ∨ x + naturals_to_semiring nat R z = 0. Proof. revert x. apply biinduction. solve_proper. exists 0. left. now rewrite preserves_0. intros n; split. intros [z [Ez | Ez]]. exists (1 + z). left. now rewrite preserves_plus, preserves_1, Ez. destruct z as [|z]. exists 1. left. rewrite O_nat_0, preserves_0, plus_0_r in Ez. rewrite Ez, preserves_1. ring. exists z. right. rewrite S_nat_1_plus, preserves_plus, preserves_1 in Ez. rewrite <-Ez. ring. intros [z [Ez | Ez]]. destruct z as [|z]. exists 1. right. rewrite O_nat_0, preserves_0 in Ez. rewrite preserves_1, <-Ez. ring. exists z. left. rewrite S_nat_1_plus, preserves_plus, preserves_1 in Ez. now apply (left_cancellation (+) 1). exists (1 + z). right. rewrite preserves_plus, preserves_1, <-Ez. ring. Qed. Lemma nat_int_nonneg_decompose x : 0 ≤ x → ∃ z, x = naturals_to_semiring nat R z. Proof. destruct (nat_int_to_semiring x) as [z [Ez1 | Ez2]]. now exists z. intros E. exists 0. rewrite preserves_0. apply (antisymmetry (≤)); auto. rewrite <-Ez2. now apply nonneg_plus_le_compat_r, to_semiring_nonneg. Qed. Lemma nat_int_le_plus x y : x ≤ y ↔ ∃ z, y = x + naturals_to_semiring nat R z. Proof. split. intros E. destruct (decompose_le E) as [z [Ez1 Ez2]]. destruct (nat_int_nonneg_decompose _ Ez1) as [u Eu]. exists u. now rewrite <-Eu. intros [z Ez]. rewrite Ez. now apply nonneg_plus_le_compat_r, to_semiring_nonneg. Qed. Lemma nat_int_lt_plus x y : x < y ↔ ∃ z, y = x + 1 + naturals_to_semiring nat R z. Proof. split. intros E. destruct (proj1 (nat_int_le_plus x y)) as [[|z] Ez]. now apply lt_le. rewrite O_nat_0, preserves_0, plus_0_r in Ez. now destruct (lt_ne_flip x y). exists z. now rewrite S_nat_1_plus, preserves_plus, preserves_1, associativity in Ez. intros [z Ez]. rewrite Ez. apply nonneg_plus_lt_compat_r. now apply to_semiring_nonneg. apply pos_plus_lt_compat_r; solve_propholds. Qed. Lemma lt_iff_plus_1_le x y : x < y ↔ x + 1 ≤ y. Proof. now rewrite nat_int_lt_plus, nat_int_le_plus. Qed. Lemma lt_iff_S_le x y : x < y ↔ 1 + x ≤ y. Proof. rewrite commutativity. now apply lt_iff_plus_1_le. Qed. Lemma pos_ge_1 x : 0 < x ↔ 1 ≤ x. Proof. split; intros E. rewrite <-(plus_0_l 1). now apply lt_iff_plus_1_le. apply lt_le_trans with 1; [solve_propholds | easy]. Qed. Lemma le_iff_lt_plus_1 x y : x ≤ y ↔ x < y + 1. Proof. split; intros E. apply lt_iff_plus_1_le. now apply (order_preserving (+1)). now apply (order_reflecting (+1)), lt_iff_plus_1_le. Qed. Lemma le_iff_lt_S x y : x ≤ y ↔ x < 1 + y. Proof. rewrite commutativity. now apply le_iff_lt_plus_1. Qed. Section another_semiring. Context `{FullPseudoSemiRingOrder R2} `{PropHolds ((1 : R2) ≶ 0)} `{!SemiRing_Morphism (f : R → R2)}. Instance: OrderPreserving f. Proof. repeat (split; try apply _). intros x y E. apply nat_int_le_plus in E. destruct E as [z E]. rewrite E, preserves_plus, (naturals.to_semiring_twice _ _ _). apply nonneg_plus_le_compat_r. now apply to_semiring_nonneg. Qed. Global Instance: StrictlyOrderPreserving f | 50. Proof. repeat (split; try apply _). intros x y E. apply nat_int_lt_plus in E. destruct E as [z E]. rewrite E, !preserves_plus, preserves_1, (naturals.to_semiring_twice _ _ _). apply nonneg_plus_lt_compat_r. now apply to_semiring_nonneg. apply pos_plus_lt_compat_r; solve_propholds. Qed. Global Instance: OrderEmbedding f | 50. Proof. split; try apply _. apply full_pseudo_order_reflecting. Qed. End another_semiring. End nat_int_order. math-classes-8.19.0/orders/naturals.v000066400000000000000000000076531460576051100175330ustar00rootroot00000000000000Require MathClasses.theory.naturals. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals MathClasses.interfaces.orders MathClasses.orders.rings. Require Export MathClasses.orders.nat_int. Section naturals_order. Context `{Naturals N} `{Apart N} `{!TrivialApart N} `{!FullPseudoSemiRingOrder Nle Nlt}. Instance nat_nonneg x : PropHolds (0 ≤ x). Proof. now apply (to_semiring_nonneg (f:=id)). Qed. Lemma nat_le_plus {x y: N}: x ≤ y ↔ ∃ z, y = x + z. Proof. split; intros E. destruct (decompose_le E) as [z [Ez1 Ez2]]. now exists z. destruct E as [z Ez]. apply compose_le with z; [solve_propholds | easy]. Qed. Lemma nat_not_neg x : ¬x < 0. Proof. apply le_not_lt_flip, nat_nonneg. Qed. Lemma nat_0_or_pos x : x = 0 ∨ 0 < x. Proof. destruct (trichotomy (<) 0 x) as [?|[?|?]]; intuition. now destruct (nat_not_neg x). Qed. Lemma nat_0_or_ge_1 x : x = 0 ∨ 1 ≤ x. Proof. rewrite <-pos_ge_1. apply nat_0_or_pos. Qed. Lemma nat_ne_0_pos x : x ≠ 0 ↔ 0 < x. Proof. split. destruct (nat_0_or_pos x); intuition. intros E1 E2. rewrite E2 in E1. now destruct (irreflexivity (<) 0). Qed. Lemma nat_ne_0_ge_1 x : x ≠ 0 ↔ 1 ≤ x. Proof. rewrite <-pos_ge_1. now apply nat_ne_0_pos. Qed. Global Instance: ∀ (z : N), PropHolds (z ≠ 0) → OrderReflecting (z *.). Proof. intros z ?. repeat (split; try apply _). apply (order_reflecting_pos (.*.) z). now apply nat_ne_0_pos. Qed. Global Program Instance slow_nat_le_dec: ∀ x y: N, Decision (x ≤ y) | 10 := λ x y, match decide (naturals_to_semiring _ nat x ≤ naturals_to_semiring _ nat y) with | left E => left _ | right E => right _ end. Next Obligation. now apply (order_reflecting (naturals_to_semiring N nat)). Qed. Section another_ring. Context `{Ring R} `{Apart R} `{!FullPseudoSemiRingOrder (A:=R) Rle Rlt} `{!SemiRing_Morphism (f : N → R)}. Lemma negate_to_ring_nonpos n : -f n ≤ 0. Proof. apply flip_nonneg_negate. now apply to_semiring_nonneg. Qed. Lemma between_to_ring n : -f n ≤ f n. Proof. apply between_nonneg. now apply to_semiring_nonneg. Qed. End another_ring. End naturals_order. #[global] Hint Extern 20 (PropHolds (_ ≤ _)) => eapply @nat_nonneg : typeclass_instances. (* A default order on the naturals *) #[global] Instance nat_le `{Naturals N} : Le N | 10 := λ x y, ∃ z, x + z = y. #[global] Instance nat_lt `{Naturals N} : Lt N | 10 := dec_lt. Section default_order. Context `{Naturals N} `{Apart N} `{!TrivialApart N}. Add Ring N2 : (rings.stdlib_semiring_theory N). Instance: Proper ((=) ==> (=) ==> iff) nat_le. Proof. intros x1 y1 E1 x2 y2 E2. split; intros [z p]; exists z. now rewrite <-E1, <-E2. now rewrite E1, E2. Qed. Instance: PartialOrder nat_le. Proof. repeat (split; try apply _). intros x. exists 0. ring. intros x y z [a A] [b B]. exists (a + b). now rewrite associativity, A, B. intros x y [a A] [b B]. destruct (naturals.zero_sum a b) as [E1 E2]. apply (left_cancellation (+) x). rewrite associativity, A, B. ring. rewrite <-A, <-B, E1, E2. ring. Qed. Instance: SemiRingOrder nat_le. Proof. repeat (split; try apply _). intros x y [a A]. now exists a. intros x y [a A]. exists a. rewrite <-A. ring. intros x y [a A]. exists a. apply (left_cancellation (+) z). rewrite <-A. ring. intros x y _ _. exists (x * y). ring. Qed. Notation n_to_sr := (naturals_to_semiring N nat). Instance: TotalRelation nat_le. Proof. assert (∀ x y, n_to_sr x ≤ n_to_sr y → x ≤ y) as P. intros x y E. destruct (decompose_le E) as [a [_ A]]. exists (naturals_to_semiring nat N a). apply (injective n_to_sr). rewrite rings.preserves_plus. now rewrite (naturals.to_semiring_involutive _ _). intros x y. destruct (total (≤) (n_to_sr x) (n_to_sr y)); [left | right]; now apply P. Qed. Global Instance: FullPseudoSemiRingOrder nat_le nat_lt. Proof. now apply dec_full_pseudo_srorder. Qed. End default_order. math-classes-8.19.0/orders/orders.v000066400000000000000000000300201460576051100171600ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.theory.strong_setoids. Lemma le_flip `{Le A} `{!TotalRelation (≤)} x y : ¬y ≤ x → x ≤ y. Proof. firstorder. Qed. Section partial_order. Context `{PartialOrder A}. Instance: Setoid A := po_setoid. Lemma eq_le x y : x = y → x ≤ y. Proof. intros E. now rewrite E. Qed. Lemma eq_le_flip x y : x = y → y ≤ x. Proof. intros E. now rewrite E. Qed. Lemma not_le_ne x y : ¬x ≤ y → x ≠ y. Proof. intros E1 E2. destruct E1. now rewrite E2. Qed. Lemma eq_iff_le x y : x = y ↔ x ≤ y ∧ y ≤ x. Proof. split; intros E. now rewrite E. now apply (antisymmetry (≤) x y). Qed. End partial_order. Section strict_order. Context `{StrictSetoidOrder A}. Instance: Setoid A := strict_setoid_order_setoid. Lemma lt_flip x y : x < y → ¬y < x. Proof. intros E1 E2. apply (irreflexivity (<) x). now transitivity y. Qed. Lemma lt_antisym x y : ¬(x < y < x). Proof. intros [E1 E2]. now destruct (lt_flip x y). Qed. Lemma lt_ne x y : x < y → x ≠ y. Proof. unfold PropHolds. intros E1 E2. rewrite E2 in E1. now destruct (irreflexivity (<) y). Qed. Lemma lt_ne_flip x y : x < y → y ≠ x. Proof. intro. now apply not_symmetry, lt_ne. Qed. Lemma eq_not_lt x y : x = y → ¬x < y. Proof. intros E. rewrite E. now apply (irreflexivity (<)). Qed. End strict_order. Section pseudo_order. Context `{PseudoOrder A}. Instance: StrongSetoid A := pseudo_order_setoid. Lemma apart_total_lt x y : x ≶ y → x < y ∨ y < x. Proof. intros. now apply apart_iff_total_lt. Qed. Lemma pseudo_order_lt_apart x y : x < y → x ≶ y. Proof. intros. apply apart_iff_total_lt. tauto. Qed. Lemma pseudo_order_lt_apart_flip x y : x < y → y ≶ x. Proof. intros. apply apart_iff_total_lt. tauto. Qed. Lemma not_lt_apart_lt_flip x y : ¬x < y → x ≶ y → y < x. Proof. rewrite apart_iff_total_lt. intuition. Qed. Lemma pseudo_order_cotrans_twice x₁ y₁ x₂ y₂ : x₁ < y₁ → x₂ < y₂ ∨ x₁ < x₂ ∨ y₂ < y₁. Proof. intros E1. destruct (cotransitive E1 x₂) as [E2|E2]; try tauto. destruct (cotransitive E2 y₂); try tauto. Qed. Lemma pseudo_order_lt_ext x₁ y₁ x₂ y₂ : x₁ < y₁ → x₂ < y₂ ∨ x₁ ≶ x₂ ∨ y₂ ≶ y₁. Proof. intros E. destruct (pseudo_order_cotrans_twice x₁ y₁ x₂ y₂ E) as [?|[?|?]]; auto using pseudo_order_lt_apart. Qed. Instance: Proper ((=) ==> (=) ==> iff) (<). Proof. assert (∀ x₁ y₁ x₂ y₂, x₁ < y₁ → x₁ = x₂ → y₁ = y₂ → x₂ < y₂) as P. intros x₁ y₁ x₂ y₂ E Ex Ey. destruct (pseudo_order_lt_ext x₁ y₁ x₂ y₂ E) as [?|[?|?]]; try tauto. contradict Ex. now apply apart_ne. contradict Ey. now apply not_symmetry, apart_ne. split; intros; eapply P; eauto; now symmetry. Qed. Global Instance: StrictSetoidOrder (_ : Lt A). Proof. repeat (split; try apply _). intros x E. destruct (pseudo_order_antisym x x); tauto. intros x y z E1 E2. destruct (cotransitive E1 z); trivial. destruct (pseudo_order_antisym y z); tauto. Qed. Global Instance: Transitive (complement (<)). Proof. intros x y z. intros E1 E2 E3. destruct (cotransitive E3 y); contradiction. Qed. Global Instance: AntiSymmetric (complement (<)). Proof. intros x y. rewrite <-tight_apart, apart_iff_total_lt. intuition. Qed. Lemma ne_total_lt `{!TrivialApart A} x y : x ≠ y → x < y ∨ y < x. Proof. rewrite <-trivial_apart. now apply apart_total_lt. Qed. Global Instance lt_trichotomy `{!TrivialApart A} `{∀ x y, Decision (x = y)} : Trichotomy (<). Proof. intros x y. destruct (decide (x = y)) as [?|?]; try tauto. destruct (ne_total_lt x y); tauto. Qed. End pseudo_order. Section full_partial_order. Context `{FullPartialOrder A}. Instance: StrongSetoid A := strict_po_setoid. (* Duplicate of strong_setoids.apart_ne. This is useful because a StrongSetoid is not defined as a substructure of a FullPartialOrder *) Instance strict_po_apart_ne x y : PropHolds (x ≶ y) → PropHolds (x ≠ y). Proof. intros. apply _. Qed. Global Instance apart_proper: Proper ((=) ==> (=) ==> iff) (≶). Proof. apply _. Qed. Global Instance: StrictSetoidOrder (<). Proof. split; try apply _. intros x1 y1 E1 x2 y2 E2. rewrite ?lt_iff_le_apart. now rewrite E1, E2. split; try apply _. intros x. red. rewrite lt_iff_le_apart. intros [_ ?]. now destruct (irreflexivity (≶) x). Qed. Lemma lt_le x y : PropHolds (x < y) → PropHolds (x ≤ y). Proof. intro. now apply lt_iff_le_apart. Qed. Lemma not_le_not_lt x y : ¬x ≤ y → ¬x < y. Proof. firstorder using lt_le. Qed. Lemma lt_apart x y : x < y → x ≶ y. Proof. intro. now apply lt_iff_le_apart. Qed. Lemma lt_apart_flip x y : x < y → y ≶ x. Proof. intro. now apply symmetry, lt_iff_le_apart. Qed. Lemma le_not_lt_flip x y : y ≤ x → ¬x < y. Proof. rewrite lt_iff_le_apart. intros E1 [E2a E2b]. contradict E2b. setoid_replace x with y. now apply (irreflexivity _). now apply (antisymmetry (≤)). Qed. Lemma lt_not_le_flip x y : y < x → ¬x ≤ y. Proof. intros E1 E2. now destruct (le_not_lt_flip y x). Qed. Lemma lt_le_trans x y z : x < y → y ≤ z → x < z. Proof. rewrite !lt_iff_le_apart. intros [E1a E1b] E2. split. now transitivity y. destruct (cotransitive E1b z) as [E3 | E3]; trivial. apply lt_apart. symmetry in E3. transitivity y; rewrite lt_iff_le_apart; tauto. Qed. Lemma le_lt_trans x y z : x ≤ y → y < z → x < z. Proof. rewrite !lt_iff_le_apart. intros E2 [E1a E1b]. split. now transitivity y. destruct (cotransitive E1b x) as [E3 | E3]; trivial. apply lt_apart. symmetry in E3. transitivity y; rewrite lt_iff_le_apart; tauto. Qed. Lemma lt_iff_le_ne `{!TrivialApart A} x y : x < y ↔ x ≤ y ∧ x ≠ y. Proof. rewrite <-trivial_apart. now apply lt_iff_le_apart. Qed. Lemma le_equiv_lt `{!TrivialApart A} `{∀ x y, Decision (x = y)} x y : x ≤ y → x = y ∨ x < y. Proof. intros. destruct (decide (x = y)); try tauto. right. rewrite lt_iff_le_ne; tauto. Qed. Program Instance dec_from_lt_dec `{!TrivialApart A} `{∀ x y, Decision (x ≤ y)} : ∀ x y, Decision (x = y) := λ x y, match decide_rel (≤) x y with | left E1 => match decide_rel (≤) y x with | left E2 => left _ | right E2 => right _ end | right E1 => right _ end. Next Obligation. now apply (antisymmetry (≤)). Qed. Next Obligation. apply not_symmetry. now apply not_le_ne. Qed. Next Obligation. now apply not_le_ne. Qed. Definition lt_dec_slow `{!TrivialApart A} `{∀ x y, Decision (x ≤ y)} : ∀ x y, Decision (x < y). Proof. intros x y. destruct (decide (x ≤ y)). destruct (decide (x = y)). right. now apply eq_not_lt. left. apply lt_iff_le_ne; intuition. right. now apply not_le_not_lt. Defined. End full_partial_order. (* Due to bug #2528 *) #[global] Hint Extern 5 (PropHolds (_ ≠ _)) => eapply @strict_po_apart_ne : typeclass_instances. #[global] Hint Extern 10 (PropHolds (_ ≤ _)) => eapply @lt_le : typeclass_instances. #[global] Hint Extern 20 (Decision (_ < _)) => eapply @lt_dec_slow : typeclass_instances. Section full_pseudo_order. Context `{FullPseudoOrder A}. Instance: StrongSetoid A := pseudo_order_setoid. Lemma not_lt_le_flip x y : ¬y < x → x ≤ y. Proof. intros. now apply le_iff_not_lt_flip. Qed. Instance: PartialOrder (≤). Proof. split; try apply _. intros ? ? E1 ? ? E2. now rewrite !le_iff_not_lt_flip, E1, E2. split. intros x. apply not_lt_le_flip, (irreflexivity (<)). intros x y z. rewrite !le_iff_not_lt_flip. intros. change (complement (<) z x). now transitivity y. intros x y. rewrite !le_iff_not_lt_flip. intros E1 E2. now apply (antisymmetry (complement (<))). Qed. Global Instance: FullPartialOrder (_ : Le A) (_ : Lt A). Proof. split; try apply _. intros x y. rewrite !le_iff_not_lt_flip. split. intros E. split. now apply lt_flip. now apply pseudo_order_lt_apart. intros [? E]. now apply not_lt_apart_lt_flip. Qed. Global Instance: ∀ x y, Stable (x ≤ y). Proof. intros x y. unfold Stable, DN. rewrite !le_iff_not_lt_flip. tauto. Qed. Lemma le_or_lt `{!TrivialApart A} `{∀ x y, Decision (x = y)} x y : x ≤ y ∨ y < x. Proof. destruct (trichotomy (<) x y) as [|[|]]; try tauto. left. now apply lt_le. left. now apply eq_le_flip. Qed. Global Instance le_total `{!TrivialApart A} `{∀ x y, Decision (x = y)} : TotalOrder (≤). Proof. split; try apply _. intros x y. destruct (le_or_lt x y). tauto. right. now apply lt_le. Qed. Lemma not_le_lt_flip `{!TrivialApart A} `{∀ x y, Decision (x = y)} x y : ¬y ≤ x → x < y. Proof. intros. destruct (le_or_lt y x); intuition. Qed. Existing Instance dec_from_lt_dec. Program Definition lt_dec `{!TrivialApart A} `{∀ x y, Decision (x ≤ y)} : ∀ x y, Decision (x < y) := λ x y, match decide_rel (≤) y x with | left E => right _ | right E => left _ end. Next Obligation. now apply le_not_lt_flip. Qed. Next Obligation. now apply not_le_lt_flip. Qed. End full_pseudo_order. #[global] Hint Extern 8 (Decision (_ < _)) => eapply @lt_dec : typeclass_instances. (* The following instances would be tempting, but turn out to be a bad idea. Hint Extern 10 (PropHolds (_ ≠ _)) => eapply @le_ne : typeclass_instances. Hint Extern 10 (PropHolds (_ ≠ _)) => eapply @le_ne_flip : typeclass_instances. It will then loop like: semirings.lt_0_1 → lt_ne_flip → ... *) Section dec_strict_setoid_order. Context `{StrictSetoidOrder A} `{Apart A} `{!TrivialApart A} `{∀ x y, Decision (x = y)}. Instance: Setoid A := strict_setoid_order_setoid. Instance: StrongSetoid A := dec_strong_setoid. Context `{!Trichotomy (<)}. Instance dec_strict_pseudo_order: PseudoOrder (<). Proof. split; try apply _. intros x y [??]. destruct (lt_antisym x y); tauto. intros x y Exy z. destruct (trichotomy (<) x z) as [? | [Exz | Exz]]; try tauto. right. now rewrite <-Exz. right. now transitivity x. intros x y. rewrite trivial_apart. split. destruct (trichotomy (<) x y); intuition. intros [?|?]. now apply lt_ne. now apply lt_ne_flip. Qed. End dec_strict_setoid_order. Section dec_partial_order. Context `{PartialOrder A} `{∀ x y, Decision (x = y)}. Instance: Setoid A := po_setoid. Definition dec_lt: Lt A := λ x y, x ≤ y ∧ x ≠ y. Context `{Alt : Lt A} (lt_correct : ∀ x y, x < y ↔ x ≤ y ∧ x ≠ y). Instance dec_order: StrictSetoidOrder (<). Proof. split; try apply _. intros ? ? E1 ? ? E2. now rewrite !lt_correct, E1, E2. split; try apply _. intros x. red. rewrite lt_correct. now intros [_ []]. intros x y z. rewrite !lt_correct. intros [E1a E1b] [E2a E2b]. split. now transitivity y. intros E3. destruct E2b. apply (antisymmetry (≤)); trivial. now rewrite <-E3. Qed. Context `{Apart A} `{!TrivialApart A}. Instance: StrongSetoid A := dec_strong_setoid. Instance dec_full_partial_order: FullPartialOrder (≤) (<). Proof. constructor; try typeclasses eauto. setoid_rewrite trivial_apart; apply lt_correct. Qed. Context `{!TotalRelation (≤)}. Instance: Trichotomy (<). Proof. intros x y. rewrite !lt_correct. destruct (decide (x = y)); try tauto. destruct (total (≤) x y); intuition. Qed. Instance dec_pseudo_order: PseudoOrder (<). Proof dec_strict_pseudo_order. Instance dec_full_pseudo_order: FullPseudoOrder (≤) (<). Proof. split; try apply _. intros x y. rewrite lt_correct. split. intros ? [? []]. now apply (antisymmetry (≤)). intros E1. destruct (total (≤) x y); trivial. destruct (decide (x = y)) as [E2|E2]. now rewrite E2. intuition. Qed. End dec_partial_order. math-classes-8.19.0/orders/rationals.v000066400000000000000000000204211460576051100176620ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring Coq.setoid_ring.Field MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.interfaces.naturals MathClasses.interfaces.rationals MathClasses.interfaces.integers MathClasses.implementations.natpair_integers MathClasses.theory.rationals MathClasses.theory.dec_fields MathClasses.theory.rings MathClasses.orders.integers MathClasses.orders.dec_fields. Section rationals_and_integers. Context `{Rationals Q} `{!SemiRingOrder Qle} Z `{Integers Z} `{Apart Z} `{!TrivialApart Z} `{!FullPseudoSemiRingOrder (A:=Z) Zle Zlt} {f : Z → Q} `{!SemiRing_Morphism f}. Add Field Q : (stdlib_field_theory Q). Lemma rationals_decompose_pos_den x : ∃ num, ∃ den, 0 < den ∧ x = f num / f den. Proof. destruct (rationals_decompose x) as [num [den [E1 E2]]]. destruct (total (≤) den 0). exists (-num), (-den). split. apply lt_iff_le_ne. split. now apply rings.flip_nonpos_negate. apply not_symmetry. now apply flip_negate_ne_0. rewrite 2!preserves_negate. rewrite E2. field. split. apply flip_negate_ne_0. now apply injective_ne_0. now apply injective_ne_0. exists num. exists den. split; try assumption. apply lt_iff_le_ne. split. assumption. now apply not_symmetry. Qed. End rationals_and_integers. (* A PseudoRingOrder uniquely specifies the orders on the rationals *) Section rationals_and_another_rationals. Context `{Rationals Q1} `{Apart Q1} `{!TrivialApart Q1} `{!FullPseudoSemiRingOrder (A:=Q1) Q1le Q1lt}. Context `{Rationals Q2} `{Apart Q2} `{!TrivialApart Q2} `{!FullPseudoSemiRingOrder (A:=Q2) Q2le Q2lt} {f : Q1 → Q2} `{!SemiRing_Morphism f}. Add Field Q1 : (stdlib_field_theory Q1). Notation i_to_r := (integers.integers_to_ring (SRpair nat) Q1). Let f_preserves_nonneg x : 0 ≤ x → 0 ≤ f x. Proof. intros E. destruct (rationals_decompose_pos_den (SRpair nat) x) as [num [den [E1 E2]]]. rewrite E2 in E |- *. clear E2. rewrite preserves_mult, preserves_dec_recip. apply (order_reflecting_pos (.*.) (f (i_to_r den))). change (0 < (f ∘ i_to_r) den). rewrite (integers.to_ring_unique _). apply semirings.preserves_pos. unfold lt in *. apply E1. apply (order_preserving_nonneg (.*.) (i_to_r den)) in E. rewrite right_absorb. rewrite right_absorb in E. rewrite (commutativity (f (i_to_r num))), associativity, dec_recip_inverse, left_identity. rewrite (commutativity (i_to_r num)), associativity, dec_recip_inverse, left_identity in E. change (0 ≤ (f ∘ i_to_r) num). rewrite (integers.to_ring_unique _). rewrite <-(preserves_0 (f:=integers_to_ring (SRpair nat) Q2)). apply (order_preserving _). apply (order_reflecting i_to_r). now rewrite preserves_0. apply injective_ne_0. now apply lt_ne_flip. change ((f ∘ i_to_r) den ≠ 0). apply injective_ne_0. now apply lt_ne_flip. apply semirings.preserves_nonneg. now apply lt_le. Qed. Instance morphism_order_preserving: OrderPreserving f. Proof. apply semirings.preserving_preserves_nonneg. apply f_preserves_nonneg. Qed. End rationals_and_another_rationals. Section rationals_order_isomorphic. Context `{Rationals Q1} `{Apart Q1} `{!TrivialApart Q1} `{!FullPseudoSemiRingOrder (A:=Q1) Q1le Q1lt} `{Rationals Q2} `{Apart Q2} `{!TrivialApart Q2} `{!FullPseudoSemiRingOrder (A:=Q2) Q2le Q2lt} {f : Q1 → Q2} `{!SemiRing_Morphism f}. Global Instance: OrderEmbedding f. Proof. split. apply morphism_order_preserving. repeat (split; try apply _). intros x y E. rewrite <-(to_rationals_involutive x (Q2:=Q2)), <-(to_rationals_involutive y (Q2:=Q2)). rewrite <-2!(to_rationals_unique f). now apply (morphism_order_preserving (f:=rationals_to_rationals Q2 Q1)). Qed. End rationals_order_isomorphic. #[global] Instance rationals_le `{Rationals Q} : Le Q | 10 := λ x y, ∃ num, ∃ den, y = x + naturals_to_semiring nat Q num / naturals_to_semiring nat Q den. #[global] Instance rationals_lt `{Rationals Q} : Lt Q | 10 := dec_lt. Section default_order. Context `{Rationals Q} `{Apart Q} `{!TrivialApart Q}. Add Field F: (stdlib_field_theory Q). Notation n_to_sr := (naturals_to_semiring nat Q). Instance: Proper ((=) ==> (=) ==> iff) rationals_le. Proof. intros x x' E y y' E'. unfold rationals_le. split; intros [n [d d_nonzero]]; exists n, d. now rewrite <-E, <-E'. now rewrite E, E'. Qed. Instance: Reflexive rationals_le. Proof. intro. exists (0%mc:nat), (0%mc:nat). rewrite preserves_0. ring. Qed. (* rationals_le can actually only happen if the denominator is nonzero: *) Lemma rationals_decompose_le (x y: Q) : x ≤ y → ∃ num, ∃ den, den ≠ 0 ∧ y = x + n_to_sr num * / n_to_sr den. Proof with eauto. intros [n [d E]]. destruct (decide (d = 0)) as [A|A]... exists (0%mc:nat), (1%mc:nat). split. discriminate. rewrite E, A, preserves_0, preserves_1, dec_recip_0. ring. Qed. Instance: Transitive rationals_le. Proof with auto. intros x y z E1 E2. destruct (rationals_decompose_le x y) as [n1 [d1 [A1 B1]]]... destruct (rationals_decompose_le y z) as [n2 [d2 [A2 B2]]]... exists (n1 * d2 + n2 * d1), (d1 * d2). rewrite B2, B1. rewrite preserves_plus. rewrite ?preserves_mult. field. split; now apply injective_ne_0. Qed. Instance: AntiSymmetric rationals_le. Proof with auto. intros x y E1 E2. destruct (rationals_decompose_le x y) as [n1 [d1 [A1 B1]]]... destruct (rationals_decompose_le y x) as [n2 [d2 [A2 B2]]]... rewrite B1 in B2 |- *. clear E1 E2 B1 y. rewrite <-associativity in B2. rewrite <-(plus_0_r x) in B2 at 1. apply (left_cancellation (+) x) in B2. destruct (zero_product n1 d2) as [F|F]... apply naturals.zero_sum with (d1 * n2). apply (injective n_to_sr). rewrite preserves_plus, preserves_mult, preserves_mult, preserves_0. apply (left_cancellation_ne_0 (.*.) (/n_to_sr d1)). apply dec_recip_ne_0. apply injective_ne_0... apply (left_cancellation_ne_0 (.*.) (/n_to_sr d2)). apply dec_recip_ne_0. apply injective_ne_0... ring_simplify. etransitivity. 2: now symmetry; eauto. field. split; apply injective_ne_0... rewrite F. rewrite preserves_0. ring. contradiction. Qed. Instance: PartialOrder rationals_le. Proof. repeat (split; try apply _). Qed. Instance: SemiRingOrder rationals_le. Proof. apply from_ring_order. repeat (split; try apply _). intros x y [n [d E]]. exists n, d. rewrite E. ring. intros x y [n1 [d1 E1]] [n2 [d2 E2]]. exists (n1 * n2), (d1 * d2). rewrite 2!preserves_mult. rewrite E1, E2, dec_recip_distr. ring. Qed. Notation i_to_r := (integers_to_ring (SRpair nat) Q). Instance: TotalRelation rationals_le. Proof with auto. assert (∀ xn xd yn yd, 0 < xd → 0 < yd → xn * yd ≤ yn * xd → i_to_r xn / i_to_r xd ≤ i_to_r yn / i_to_r yd) as P. intros xn xd yn yd. rewrite !lt_iff_le_apart. intros [xd_ge0 xd_ne0] [yd_ge0 yd_ne0] E. destruct (semirings.decompose_le E) as [z [Ez1 Ex2]]. apply nat_int_le_plus in xd_ge0. apply nat_int_le_plus in yd_ge0. apply nat_int_le_plus in Ez1. destruct xd_ge0 as [xd' xd_ge0], yd_ge0 as [yd' yd_ge0], Ez1 as [z' Ez1]. rewrite left_identity in xd_ge0, yd_ge0, Ez1. exists z'. exists (xd' * yd'). assert (∀ a, (i_to_r ∘ naturals_to_semiring nat (SRpair nat)) a = n_to_sr a) as F. intros a. apply (naturals.to_semiring_unique _). rewrite preserves_mult, <-F, <-F, <-F. unfold compose. rewrite <-xd_ge0, <-yd_ge0, <-Ez1. transitivity ((i_to_r yn * i_to_r xd) / (i_to_r yd * i_to_r xd)). field. split; apply injective_ne_0; apply not_symmetry... rewrite <-preserves_mult, Ex2, preserves_plus, preserves_mult. field. split; apply injective_ne_0; apply not_symmetry... intros x y. destruct (rationals_decompose_pos_den (SRpair nat) x) as [xn [xd [E1x E2x]]]. destruct (rationals_decompose_pos_den (SRpair nat) y) as [yn [yd [E1y E2y]]]. rewrite E2x, E2y. destruct (total (≤) (xn * yd) (yn * xd)); [left | right]; now apply P. Qed. Global Instance: FullPseudoSemiRingOrder rationals_le rationals_lt. Proof. now apply dec_full_pseudo_srorder. Qed. End default_order. math-classes-8.19.0/orders/rings.v000066400000000000000000000241771460576051100170240ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.theory.rings. Require Export MathClasses.orders.semirings. Section from_ring_order. Context `{Ring R} `{!PartialOrder Rle} (plus_spec : ∀ z, OrderPreserving (z +)) (mult_spec : ∀ x y, PropHolds (0 ≤ x) → PropHolds (0 ≤ y) → PropHolds (0 ≤ x * y)). Lemma from_ring_order: SemiRingOrder (≤). Proof. repeat (split; try apply _). intros x y E. exists (- x + y). now rewrite associativity, plus_negate_r, plus_0_l. intros x y E. rewrite <-(plus_0_l x), <-(plus_0_l y), <-!(plus_negate_l z), <-!associativity. now apply (order_preserving _). Qed. End from_ring_order. Section from_strict_ring_order. Context `{Ring R} `{!StrictSetoidOrder Rle} (plus_spec : ∀ z, StrictlyOrderPreserving (z +)) (mult_spec : ∀ x y, PropHolds (0 < x) → PropHolds (0 < y) → PropHolds (0 < x * y)). Lemma from_strict_ring_order: StrictSemiRingOrder (<). Proof. repeat (split; try apply _). intros x y E. exists (- x + y). now rewrite associativity, plus_negate_r, plus_0_l. intros x y E. rewrite <-(plus_0_l x), <-(plus_0_l y), <-!(plus_negate_l z), <-!associativity. now apply (strictly_order_preserving _). Qed. End from_strict_ring_order. Section from_pseudo_ring_order. Context `{Ring R} `{Apart R} `{!PseudoOrder Rlt} (plus_spec : ∀ z, StrictlyOrderPreserving (z +)) (mult_ext : StrongSetoid_BinaryMorphism (.*.)) (mult_spec : ∀ x y, PropHolds (0 < x) → PropHolds (0 < y) → PropHolds (0 < x * y)). Lemma from_pseudo_ring_order: PseudoSemiRingOrder (<). Proof. repeat (split; try apply _). intros x y E. exists (- x + y). now rewrite associativity, plus_negate_r, plus_0_l. intros x y E. rewrite <-(plus_0_l x), <-(plus_0_l y), <-!(plus_negate_l z), <-!associativity. now apply (strictly_order_preserving _). Qed. End from_pseudo_ring_order. Section from_full_pseudo_ring_order. Context `{Ring R} `{Apart R} `{!FullPseudoOrder Rle Rlt} (plus_spec : ∀ z, StrictlyOrderPreserving (z +)) (mult_ext : StrongSetoid_BinaryMorphism (.*.)) (mult_spec : ∀ x y, PropHolds (0 < x) → PropHolds (0 < y) → PropHolds (0 < x * y)). Lemma from_full_pseudo_ring_order: FullPseudoSemiRingOrder (≤) (<). Proof. split. now apply from_pseudo_ring_order. now apply le_iff_not_lt_flip. Qed. End from_full_pseudo_ring_order. Section ring_order. Context `{Ring R} `{!SemiRingOrder Rle}. Add Ring R : (stdlib_ring_theory R). Lemma flip_le_negate x y : -y ≤ -x ↔ x ≤ y. Proof. assert (∀ a b, a ≤ b → -b ≤ -a). intros a b E. setoid_replace (-b) with (-a + -b + a) by ring. setoid_replace (-a) with (-a + -b + b) at 2 by ring. now apply (order_preserving _). split; intros; auto. rewrite <-(negate_involutive x), <-(negate_involutive y); auto. Qed. Lemma flip_nonneg_negate x : 0 ≤ x ↔ -x ≤ 0. Proof. split; intros E. rewrite <-negate_0. now apply flip_le_negate. apply flip_le_negate. now rewrite negate_0. Qed. Lemma flip_nonpos_negate x : x ≤ 0 ↔ 0 ≤ -x. Proof. rewrite <-(negate_involutive x) at 1. split; intros; now apply flip_nonneg_negate. Qed. Lemma flip_le_minus_r (x y z : R) : z ≤ y - x ↔ z + x ≤ y. Proof. split; intros E. rewrite commutativity. setoid_replace y with (x + (y - x)) by ring. now apply (order_preserving _). rewrite commutativity. setoid_replace z with (-x + (z + x)) by ring. now apply (order_preserving _). Qed. Lemma flip_le_minus_l (x y z : R) : y - x ≤ z ↔ y ≤ z + x. Proof. rewrite <-(negate_involutive x) at 2. split; now apply flip_le_minus_r. Qed. Lemma flip_nonneg_minus (x y : R) : 0 ≤ y - x ↔ x ≤ y. Proof. setoid_replace x with (0 + x) at 2 by ring. now apply flip_le_minus_r. Qed. Lemma flip_nonpos_minus (x y : R) : y - x ≤ 0 ↔ y ≤ x. Proof. setoid_replace x with (0 + x) at 2 by ring. now apply flip_le_minus_l. Qed. Lemma nonneg_minus_compat (x y z : R) : 0 ≤ z → x ≤ y → x - z ≤ y. Proof. intros E1 E2. rewrite commutativity. setoid_replace y with (-z + (y + z)) by ring. apply (order_preserving (-(z) +)). transitivity y; trivial. setoid_replace y with (y + 0) at 1 by ring. now apply (order_preserving (y +)). Qed. Lemma nonneg_minus_compat_back (x y z : R) : 0 ≤ z → x ≤ y - z → x ≤ y. Proof. intros E1 E2. transitivity (y - z); trivial. now apply nonneg_minus_compat. Qed. Lemma between_nonneg (x : R) : 0 ≤ x → -x ≤ x. Proof. intros E. transitivity 0; trivial. now apply flip_nonneg_negate. Qed. End ring_order. Section strict_ring_order. Context `{Ring R} `{!StrictSemiRingOrder Rlt}. Add Ring Rs : (stdlib_ring_theory R). Lemma flip_lt_negate x y : -y < -x ↔ x < y. Proof. assert (∀ a b, a < b → -b < -a). intros a b E. setoid_replace (-b) with (-a + -b + a) by ring. setoid_replace (-a) with (-a + -b + b) at 2 by ring. now apply (strictly_order_preserving _). split; intros; auto. rewrite <-(negate_involutive x), <-(negate_involutive y); auto. Qed. Lemma flip_pos_negate x : 0 < x ↔ -x < 0. Proof. split; intros E. rewrite <- negate_0. now apply flip_lt_negate. apply flip_lt_negate. now rewrite negate_0. Qed. Lemma flip_neg_negate x : x < 0 ↔ 0 < -x. Proof. rewrite <-(negate_involutive x) at 1. split; intros; now apply flip_pos_negate. Qed. Lemma flip_lt_minus_r (x y z : R) : z < y - x ↔ z + x < y. Proof. split; intros E. rewrite commutativity. setoid_replace y with (x + (y - x)) by ring. now apply (strictly_order_preserving _). rewrite commutativity. setoid_replace z with (-x + (z + x)) by ring. now apply (strictly_order_preserving _). Qed. Lemma flip_lt_minus_l (x y z : R) : y - x < z ↔ y < z + x. Proof. rewrite <-(negate_involutive x) at 2. split; now apply flip_lt_minus_r. Qed. Lemma flip_pos_minus (x y : R) : 0 < y - x ↔ x < y. Proof. setoid_replace x with (0 + x) at 2 by ring. now apply flip_lt_minus_r. Qed. Lemma flip_neg_minus (x y : R) : y - x < 0 ↔ y < x. Proof. setoid_replace x with (0 + x) at 2 by ring. now apply flip_lt_minus_l. Qed. Lemma pos_minus_compat (x y z : R) : 0 < z → x < y → x - z < y. Proof. intros E1 E2. rewrite commutativity. setoid_replace y with (-z + (y + z)) by ring. apply (strictly_order_preserving (-(z) +)). transitivity y; trivial. setoid_replace y with (y + 0) at 1 by ring. now apply (strictly_order_preserving (y +)). Qed. Lemma between_pos (x : R) : 0 < x → -x < x. Proof. intros E. transitivity 0; trivial. now apply flip_pos_negate. Qed. End strict_ring_order. Section another_ring_order. Context `{Ring R1} `{!SemiRingOrder R1le} `{Ring R2} `{R2le : Le R2}. Lemma projected_ring_order (f : R2 → R1) `{!SemiRing_Morphism f} `{!Injective f} : (∀ x y, x ≤ y ↔ f x ≤ f y) → SemiRingOrder R2le. Proof. intros P. apply (projected_srorder f P). intros x y E. exists (-x + y). now rewrite associativity, plus_negate_r, plus_0_l. Qed. Context `{!SemiRingOrder R2le} {f : R1 → R2} `{!SemiRing_Morphism f}. Lemma reflecting_preserves_nonneg : (∀ x, 0 ≤ f x → 0 ≤ x) → OrderReflecting f. Proof. intros E. repeat (split; try apply _). intros x y F. apply flip_nonneg_minus, E. rewrite preserves_plus, preserves_negate. now apply flip_nonneg_minus, F. Qed. Lemma preserves_ge_negate1 `{!OrderPreserving f} x : -1 ≤ x → -1 ≤ f x. Proof. intros. rewrite <-(preserves_1 (f:=f)), <-preserves_negate. now apply (order_preserving f). Qed. Lemma preserves_le_negate1 `{!OrderPreserving f} x : x ≤ -1 → f x ≤ -1. Proof. intros. rewrite <-(preserves_1 (f:=f)), <-preserves_negate. now apply (order_preserving f). Qed. End another_ring_order. Section another_strict_ring_order. Context `{Ring R1} `{!StrictSemiRingOrder R1lt} `{Ring R2} `{R2lt : Lt R2}. Lemma projected_strict_ring_order (f : R2 → R1) `{!SemiRing_Morphism f} : (∀ x y, x < y ↔ f x < f y) → StrictSemiRingOrder R2lt. Proof. intros P. pose proof (projected_strict_setoid_order f P). apply from_strict_ring_order. repeat (split; try apply _). intros x y E. apply P. rewrite 2!preserves_plus. now apply (strictly_order_preserving _), P. intros x y E1 E2. apply P. rewrite preserves_mult, preserves_0. now apply pos_mult_compat; rewrite <-(preserves_0 (f:=f)); apply P. Qed. End another_strict_ring_order. Section another_pseudo_ring_order. Context `{Ring R1} `{Apart R1} `{!PseudoSemiRingOrder R1lt} `{Ring R2} `{Apart R2} `{R2lt : Lt R2}. Lemma projected_pseudo_ring_order (f : R2 → R1) `{!SemiRing_Morphism f} `{!StrongInjective f} : (∀ x y, x < y ↔ f x < f y) → PseudoSemiRingOrder R2lt. Proof. intros P. pose proof (projected_pseudo_order f P). pose proof (projected_strict_ring_order f P). apply from_pseudo_ring_order; try apply _. pose proof (pseudo_order_setoid : StrongSetoid R1). pose proof (pseudo_order_setoid : StrongSetoid R2). pose proof (strong_injective_mor f). repeat (split; try apply _). intros x₁ y₁ x₂ y₂ E. apply (strong_injective f) in E. rewrite 2!preserves_mult in E. destruct (strong_binary_extensionality (.*.) _ _ _ _ E); [left | right]; now apply (strong_extensionality f). Qed. End another_pseudo_ring_order. Section another_full_pseudo_ring_order. Context `{Ring R1} `{Apart R1} `{!FullPseudoSemiRingOrder R1le R1lt} `{Ring R2} `{Apart R2} `{R2le : Le R2} `{R2lt : Lt R2}. Lemma projected_full_pseudo_ring_order (f : R2 → R1) `{!SemiRing_Morphism f} `{!StrongInjective f} : (∀ x y, x ≤ y ↔ f x ≤ f y) → (∀ x y, x < y ↔ f x < f y) → FullPseudoSemiRingOrder R2le R2lt. Proof. intros P1 P2. pose proof (projected_full_pseudo_order f P1 P2). pose proof (projected_pseudo_ring_order f P2). split; try apply _. apply le_iff_not_lt_flip. Qed. End another_full_pseudo_ring_order. math-classes-8.19.0/orders/semirings.v000066400000000000000000000575641460576051100177100ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.theory.rings. Require Export MathClasses.orders.orders MathClasses.orders.maps. Local Existing Instance srorder_semiring. Local Existing Instance strict_srorder_semiring. Local Existing Instance pseudo_srorder_semiring. Section semiring_order. Context `{SemiRingOrder R}. Add Ring R : (stdlib_semiring_theory R). Global Instance: ∀ (z : R), OrderEmbedding (+z). Proof. intro. split. now apply order_preserving_flip. now apply order_reflecting_flip. Qed. Global Instance: ∀ z, LeftCancellation (+) z. Proof. intros z x y E. apply (antisymmetry (≤)); apply (order_reflecting (z+)); now apply eq_le. Qed. Global Instance: ∀ z, RightCancellation (+) z. Proof. intros. apply (right_cancel_from_left (+)). Qed. Lemma nonneg_plus_le_compat_r x z : 0 ≤ z ↔ x ≤ x + z. Proof. rewrite <-(plus_0_r x) at 1. split; intros. now apply (order_preserving _). now apply (order_reflecting (x+)). Qed. Lemma nonneg_plus_le_compat_l x z : 0 ≤ z ↔ x ≤ z + x. Proof. rewrite commutativity. now apply nonneg_plus_le_compat_r. Qed. Lemma plus_le_compat x₁ y₁ x₂ y₂ : x₁ ≤ y₁ → x₂ ≤ y₂ → x₁ + x₂ ≤ y₁ + y₂. Proof. intros E1 E2. transitivity (y₁ + x₂). now apply (order_preserving (+ x₂)). now apply (order_preserving (y₁ +)). Qed. Lemma plus_le_compat_r x y z : 0 ≤ z → x ≤ y → x ≤ y + z. Proof. intros. rewrite <-(plus_0_r x). now apply plus_le_compat. Qed. Lemma plus_le_compat_l x y z : 0 ≤ z → x ≤ y → x ≤ z + y. Proof. rewrite commutativity. now apply plus_le_compat_r. Qed. Lemma nonpos_plus_compat x y : x ≤ 0 → y ≤ 0 → x + y ≤ 0. Proof. intros. rewrite <-(plus_0_r 0). now apply plus_le_compat. Qed. Instance nonneg_plus_compat (x y : R) : PropHolds (0 ≤ x) → PropHolds (0 ≤ y) → PropHolds (0 ≤ x + y). Proof. intros. now apply plus_le_compat_l. Qed. Lemma decompose_le {x y} : x ≤ y → ∃ z, 0 ≤ z ∧ y = x + z. Proof. intros E. destruct (srorder_partial_minus x y E) as [z Ez]. exists z. split; [| easy]. apply (order_reflecting (x+)). now rewrite plus_0_r, <-Ez. Qed. Lemma compose_le x y z : 0 ≤ z → y = x + z → x ≤ y. Proof. intros E1 E2. rewrite E2. now apply nonneg_plus_le_compat_r. Qed. Global Instance: ∀ (z : R), PropHolds (0 ≤ z) → OrderPreserving (z *.). Proof. intros z E. repeat (split; try apply _). intros x y F. destruct (decompose_le F) as [a [Ea1 Ea2]]. rewrite Ea2, plus_mult_distr_l. apply nonneg_plus_le_compat_r. now apply nonneg_mult_compat. Qed. Global Instance: ∀ (z : R), PropHolds (0 ≤ z) → OrderPreserving (.* z). Proof. intros. now apply order_preserving_flip. Qed. Lemma mult_le_compat x₁ y₁ x₂ y₂ : 0 ≤ x₁ → 0 ≤ x₂ → x₁ ≤ y₁ → x₂ ≤ y₂ → x₁ * x₂ ≤ y₁ * y₂. Proof. intros Ex₁ Ey₁ E1 E2. transitivity (y₁ * x₂). now apply (order_preserving_flip_nonneg (.*.) x₂). apply (order_preserving_nonneg (.*.) y₁); [| easy]. now transitivity x₁. Qed. Lemma ge_1_mult_le_compat_r x y z : 1 ≤ z → 0 ≤ y → x ≤ y → x ≤ y * z. Proof. intros. transitivity y; [easy |]. rewrite <-(mult_1_r y) at 1. now apply (order_preserving_nonneg (.*.) y). Qed. Lemma ge_1_mult_le_compat_l x y z : 1 ≤ z → 0 ≤ y → x ≤ y → x ≤ z * y. Proof. rewrite commutativity. now apply ge_1_mult_le_compat_r. Qed. Lemma flip_nonpos_mult_l x y z : z ≤ 0 → x ≤ y → z * y ≤ z * x. Proof. intros Ez Exy. destruct (decompose_le Ez) as [a [Ea1 Ea2]], (decompose_le Exy) as [b [Eb1 Eb2]]. rewrite Eb2. apply compose_le with (a * b). now apply nonneg_mult_compat. transitivity (z * x + (z + a) * b). rewrite <-Ea2. ring. ring. Qed. Lemma flip_nonpos_mult_r x y z : z ≤ 0 → x ≤ y → y * z ≤ x * z. Proof. rewrite 2!(commutativity _ z). now apply flip_nonpos_mult_l. Qed. Lemma nonpos_mult x y : x ≤ 0 → y ≤ 0 → 0 ≤ x * y. Proof. intros. rewrite <-(mult_0_r x). now apply flip_nonpos_mult_l. Qed. Lemma nonpos_nonneg_mult x y : x ≤ 0 → 0 ≤ y → x * y ≤ 0. Proof. intros. rewrite <-(mult_0_r x). now apply flip_nonpos_mult_l. Qed. Lemma nonneg_nonpos_mult x y : 0 ≤ x → y ≤ 0 → x * y ≤ 0. Proof. intros. rewrite commutativity. now apply nonpos_nonneg_mult. Qed. End semiring_order. (* Due to bug #2528 *) #[global] Hint Extern 7 (PropHolds (0 ≤ _ + _)) => eapply @nonneg_plus_compat : typeclass_instances. Section strict_semiring_order. Context `{SemiRing R} `{Apart R} `{!StrictSemiRingOrder Rlt}. Add Ring Rs : (stdlib_semiring_theory R). Global Instance: ∀ (z : R), StrictOrderEmbedding (+z). Proof. intro. split. now apply strictly_order_preserving_flip. now apply strictly_order_reflecting_flip. Qed. Lemma pos_plus_lt_compat_r x z : 0 < z ↔ x < x + z. Proof. rewrite <-(plus_0_r x) at 1. split; intros. now apply (strictly_order_preserving _). now apply (strictly_order_reflecting (x+)). Qed. Lemma pos_plus_lt_compat_l x z : 0 < z → x < z + x. Proof. rewrite commutativity. now apply pos_plus_lt_compat_r. Qed. Lemma plus_lt_compat x₁ y₁ x₂ y₂ : x₁ < y₁ → x₂ < y₂ → x₁ + x₂ < y₁ + y₂. Proof. intros E1 E2. transitivity (y₁ + x₂). now apply (strictly_order_preserving (+ x₂)). now apply (strictly_order_preserving (y₁ +)). Qed. Lemma plus_lt_compat_r x y z : 0 < z → x < y → x < y + z. Proof. intros. rewrite <-(plus_0_r x). now apply plus_lt_compat. Qed. Lemma plus_lt_compat_l x y z : 0 < z → x < y → x < z + y. Proof. rewrite commutativity. now apply plus_lt_compat_r. Qed. Lemma neg_plus_compat x y : x < 0 → y < 0 → x + y < 0. Proof. intros. rewrite <-(plus_0_r 0). now apply plus_lt_compat. Qed. Instance pos_plus_compat (x y : R) : PropHolds (0 < x) → PropHolds (0 < y) → PropHolds (0 < x + y). Proof. intros. now apply plus_lt_compat_l. Qed. Lemma compose_lt x y z : 0 < z → y = x + z → x < y. Proof. intros E1 E2. rewrite E2. now apply pos_plus_lt_compat_r. Qed. Lemma decompose_lt {x y} : x < y → ∃ z, 0 < z ∧ y = x + z. Proof. intros E. destruct (strict_srorder_partial_minus x y E) as [z Ez]. exists z. split; [| easy]. apply (strictly_order_reflecting (x+)). now rewrite <-Ez, rings.plus_0_r. Qed. Global Instance: ∀ (z : R), PropHolds (0 < z) → StrictlyOrderPreserving (z *.). Proof. intros z E. repeat (split; try apply _). intros x y F. destruct (decompose_lt F) as [a [Ea1 Ea2]]. rewrite Ea2, plus_mult_distr_l. apply pos_plus_lt_compat_r. now apply pos_mult_compat. Qed. Global Instance: ∀ (z : R), PropHolds (0 < z) → StrictlyOrderPreserving (.* z). Proof. intros. now apply strictly_order_preserving_flip. Qed. Lemma mult_lt_compat x₁ y₁ x₂ y₂ : 0 < x₁ → 0 < x₂ → x₁ < y₁ → x₂ < y₂ → x₁ * x₂ < y₁ * y₂. Proof. intros Ex₁ Ey₁ E1 E2. transitivity (y₁ * x₂). now apply (strictly_order_preserving_flip_pos (.*.) x₂). apply (strictly_order_preserving_pos (.*.) y₁); [| easy ]. now transitivity x₁. Qed. Lemma gt_1_mult_lt_compat_r x y z : 1 < z → 0 < y → x < y → x < y * z. Proof. intros. transitivity y; [ easy |]. rewrite <-(mult_1_r y) at 1. now apply (strictly_order_preserving_pos (.*.) y). Qed. Lemma gt_1_mult_lt_compat_l x y z : 1 < z → 0 < y → x < y → x < z * y. Proof. rewrite commutativity. now apply gt_1_mult_lt_compat_r. Qed. Lemma flip_neg_mult_l x y z : z < 0 → x < y → z * y < z * x. Proof. intros Ez Exy. destruct (decompose_lt Ez) as [a [Ea1 Ea2]], (decompose_lt Exy) as [b [Eb1 Eb2]]. rewrite Eb2. apply compose_lt with (a * b). now apply pos_mult_compat. transitivity (z * x + (z + a) * b). rewrite <-Ea2. ring. ring. Qed. Lemma flip_neg_mult_r x y z : z < 0 → x < y → y * z < x * z. Proof. rewrite 2!(commutativity _ z). now apply flip_neg_mult_l. Qed. Lemma neg_mult x y : x < 0 → y < 0 → 0 < x * y. Proof. intros. rewrite <-(mult_0_r x). now apply flip_neg_mult_l. Qed. Lemma neg_pos_mult x y : x < 0 → 0 < y → x * y < 0. Proof. intros. rewrite <-(mult_0_r x). now apply flip_neg_mult_l. Qed. Lemma pos_neg_mult x y : 0 < x → y < 0 → x * y < 0. Proof. intros. rewrite commutativity. now apply neg_pos_mult. Qed. End strict_semiring_order. (* Due to bug #2528 *) #[global] Hint Extern 7 (PropHolds (0 < _ + _)) => eapply @pos_plus_compat : typeclass_instances. Section pseudo_semiring_order. Context `{PseudoSemiRingOrder R}. Add Ring Rp : (stdlib_semiring_theory R). Instance: StrongSetoid R := pseudo_order_setoid. Global Instance: StrictSemiRingOrder (_ : Lt R). Proof. split; try apply _. intros. now apply pseudo_srorder_partial_minus, lt_flip. now apply pseudo_srorder_pos_mult_compat. Qed. Global Instance: StrongSetoid_BinaryMorphism (+). Proof. assert (∀ z, StrongSetoid_Morphism (z +)). intros. apply pseudo_order_embedding_ext. now apply strong_setoids.strong_binary_setoid_morphism_commutative. Qed. Global Instance: ∀ z, StrongLeftCancellation (+) z. Proof. intros z x y. rewrite !apart_iff_total_lt. intros [|]; [left | right]; now apply (strictly_order_preserving (z +)). Qed. Global Instance: ∀ z, StrongRightCancellation (+) z. Proof. intros. now apply (strong_right_cancel_from_left (+)). Qed. Lemma neg_mult_decompose x y : x * y < 0 → (x < 0 ∧ 0 < y) ∨ (0 < x ∧ y < 0). Proof. intros. assert (0 ≶ x) as Ex. apply (strong_extensionality (.* y)). rewrite mult_0_l. now apply pseudo_order_lt_apart_flip. assert (0 ≶ y) as Ey. apply (strong_extensionality (x *.)). rewrite mult_0_r. now apply pseudo_order_lt_apart_flip. rewrite apart_iff_total_lt in Ex, Ey. destruct Ex as [Ex|Ex], Ey as [Ey|Ey]; try tauto. destruct (irreflexivity (<) 0). transitivity (x * y); [| easy]. now apply pos_mult_compat. destruct (irreflexivity (<) 0). transitivity (x * y); [| easy]. now apply neg_mult. Qed. Lemma pos_mult_decompose x y : 0 < x * y → (0 < x ∧ 0 < y) ∨ (x < 0 ∧ y < 0). Proof. intros. assert (0 ≶ x) as Ex. apply (strong_extensionality (.* y)). rewrite mult_0_l. now apply pseudo_order_lt_apart. assert (0 ≶ y) as Ey. apply (strong_extensionality (x *.)). rewrite mult_0_r. now apply pseudo_order_lt_apart. rewrite apart_iff_total_lt in Ex, Ey. destruct Ex as [Ex|Ex], Ey as [Ey|Ey]; try tauto. destruct (irreflexivity (<) 0). transitivity (x * y); [easy |]. now apply pos_neg_mult. destruct (irreflexivity (<) 0). transitivity (x * y); [easy |]. now apply neg_pos_mult. Qed. Global Instance: ∀ (z : R), PropHolds (0 < z) → StrictlyOrderReflecting (z *.). Proof. intros z Ez. repeat (split; try apply _). intros x y E1. apply not_lt_apart_lt_flip. intros E2. apply (lt_flip _ _ E1). now apply (strictly_order_preserving (z *.)). apply (strong_extensionality (z *.)). now apply pseudo_order_lt_apart_flip. Qed. Global Instance: ∀ (z : R), PropHolds (0 < z) → StrictlyOrderReflecting (.* z). Proof. intros. now apply strictly_order_reflecting_flip. Qed. Global Instance: ∀ z, PropHolds (z ≶ 0) → StrongLeftCancellation (.*.) z. Proof. intros z Ez x y E. red in Ez. rewrite apart_iff_total_lt in E, Ez |- *. destruct E as [E|E], Ez as [Ez|Ez]. right. now apply flip_neg_mult_l. left. now apply (strictly_order_preserving_pos (.*.) z). left. now apply flip_neg_mult_l. right. now apply (strictly_order_preserving_pos (.*.) z). Qed. Global Instance: ∀ z, PropHolds (z ≶ 0) → StrongRightCancellation (.*.) z. Proof. intros. now apply (strong_right_cancel_from_left (.*.)). Qed. Global Instance: ∀ z, PropHolds (z ≶ 0) → LeftCancellation (.*.) z. Proof. intros. now apply _. Qed. Global Instance: ∀ z, PropHolds (z ≶ 0) → RightCancellation (.*.) z. Proof. intros. now apply _. Qed. Lemma square_pos x : x ≶ 0 → 0 < x * x. Proof. intros E. apply apart_iff_total_lt in E. destruct E as [E|E]. destruct (decompose_lt E) as [z [Ez1 Ez2]]. apply compose_lt with (z * z). now apply pos_mult_compat. rewrite plus_0_l. apply (left_cancellation (+) (x * z)). rewrite <-plus_mult_distr_r, <-plus_mult_distr_l. rewrite (commutativity z x), <-!Ez2. ring. now apply pos_mult_compat. Qed. Lemma pos_mult_rev_l x y : 0 < x * y → 0 < y → 0 < x. Proof. intros. apply (strictly_order_reflecting (.* y)). now rewrite rings.mult_0_l. Qed. Lemma pos_mult_rev_r x y : 0 < x * y → 0 < x → 0 < y. Proof. intros. apply pos_mult_rev_l with x. now rewrite commutativity. easy. Qed. Context `{PropHolds (1 ≶ 0)}. Instance lt_0_1 : PropHolds (0 < 1). Proof. red. setoid_replace 1 with (1 * 1) by ring. now apply square_pos. Qed. Instance lt_0_2 : PropHolds (0 < 2). Proof. apply _. Qed. Instance lt_0_3 : PropHolds (0 < 3). Proof. apply _. Qed. Instance lt_0_4 : PropHolds (0 < 4). Proof. apply _. Qed. Lemma lt_1_2 : 1 < 2. Proof. now apply pos_plus_lt_compat_r, lt_0_1. Qed. Lemma lt_1_3 : 1 < 3. Proof. now apply pos_plus_lt_compat_r, lt_0_2. Qed. Lemma lt_1_4 : 1 < 4. Proof. now apply pos_plus_lt_compat_r, lt_0_3. Qed. Lemma lt_2_3 : 2 < 3. Proof. apply (strictly_order_preserving (1+)), lt_1_2. Qed. Lemma lt_2_4 : 2 < 4. Proof. apply (strictly_order_preserving (1+)), lt_1_3. Qed. Lemma lt_3_4 : 3 < 4. Proof. apply (strictly_order_preserving (1+)), lt_2_3. Qed. Instance apart_0_2 : PropHolds (2 ≶ 0). Proof. red. symmetry. now apply pseudo_order_lt_apart, lt_0_2. Qed. End pseudo_semiring_order. #[global] Hint Extern 7 (PropHolds (0 < 1)) => eapply @lt_0_1 : typeclass_instances. #[global] Hint Extern 7 (PropHolds (0 < 2)) => eapply @lt_0_2 : typeclass_instances. #[global] Hint Extern 7 (PropHolds (0 < 3)) => eapply @lt_0_3 : typeclass_instances. #[global] Hint Extern 7 (PropHolds (0 < 4)) => eapply @lt_0_4 : typeclass_instances. #[global] Hint Extern 7 (PropHolds (2 ≶ 0)) => eapply @apart_0_2 : typeclass_instances. Section full_pseudo_semiring_order. Context `{FullPseudoSemiRingOrder R}. Add Ring Rf : (stdlib_semiring_theory R). Global Instance: FullPseudoOrder (_ : Le R) (_ : Lt R). Proof. split. apply _. apply full_pseudo_srorder_le_iff_not_lt_flip. Qed. Global Instance: SemiRingOrder (_ : Le R). Proof. split; try apply _. intros x y. rewrite le_iff_not_lt_flip. now apply pseudo_srorder_partial_minus. intros z. repeat (split; try apply _); intros x y. rewrite !le_iff_not_lt_flip. intros E1 E2. apply E1. now apply (strictly_order_reflecting (z+)). rewrite !le_iff_not_lt_flip. intros E1 E2. apply E1. now apply (strictly_order_preserving _). intros x y. rewrite !le_iff_not_lt_flip. intros Ex Ey E. destruct (neg_mult_decompose x y); now intuition. Qed. Global Instance: ∀ (z : R), PropHolds (0 < z) → OrderReflecting (z *.). Proof. intros z E. now apply full_pseudo_order_reflecting. Qed. Global Instance: ∀ (z : R), PropHolds (0 < z) → OrderReflecting (.* z). Proof. intros. now apply order_reflecting_flip. Qed. Lemma plus_lt_le_compat x₁ y₁ x₂ y₂ : x₁ < y₁ → x₂ ≤ y₂ → x₁ + x₂ < y₁ + y₂. Proof. intros E1 E2. apply lt_le_trans with (y₁ + x₂). now apply (strictly_order_preserving (+ x₂)). now apply (order_preserving (y₁ +)). Qed. Lemma plus_le_lt_compat x₁ y₁ x₂ y₂ : x₁ ≤ y₁ → x₂ < y₂ → x₁ + x₂ < y₁ + y₂. Proof. intros E1 E2. apply le_lt_trans with (y₁ + x₂). now apply (order_preserving (+ x₂)). now apply (strictly_order_preserving (y₁ +)). Qed. Lemma nonneg_plus_lt_compat_r x y z : 0 ≤ z → x < y → x < y + z. Proof. intros. rewrite <-(plus_0_r x). now apply plus_lt_le_compat. Qed. Lemma nonneg_plus_lt_compat_l x y z : 0 ≤ z → x < y → x < z + y. Proof. intros. rewrite commutativity. now apply nonneg_plus_lt_compat_r. Qed. Lemma pos_plus_le_lt_compat_r x y z : 0 < z → x ≤ y → x < y + z. Proof. intros. rewrite <-(plus_0_r x). now apply plus_le_lt_compat. Qed. Lemma pos_plus_le_lt_compat_l x y z : 0 < z → x ≤ y → x < z + y. Proof. intros. rewrite commutativity. now apply pos_plus_le_lt_compat_r. Qed. Lemma square_nonneg x : 0 ≤ x * x. Proof. apply not_lt_le_flip. intros E. destruct (lt_antisym (x * x) 0). split; [easy |]. apply square_pos. apply (strong_extensionality (x *.)). rewrite mult_0_r. now apply lt_apart. Qed. Lemma nonneg_mult_rev_l x y : 0 ≤ x * y → 0 < y → 0 ≤ x. Proof. intros. apply (order_reflecting (.* y)). now rewrite rings.mult_0_l. Qed. Lemma nonneg_mult_rev_r x y : 0 ≤ x * y → 0 < x → 0 ≤ y. Proof. intros. apply nonneg_mult_rev_l with x. now rewrite commutativity. easy. Qed. Instance le_0_1 : PropHolds (0 ≤ 1). Proof. red. setoid_replace 1 with (1 * 1) by ring. now apply square_nonneg. Qed. Instance le_0_2 : PropHolds (0 ≤ 2). Proof. solve_propholds. Qed. Instance le_0_3 : PropHolds (0 ≤ 3). Proof. solve_propholds. Qed. Instance le_0_4 : PropHolds (0 ≤ 4). Proof. solve_propholds. Qed. Lemma le_1_2 : 1 ≤ 2. Proof. now apply nonneg_plus_le_compat_r, le_0_1. Qed. Lemma le_1_3 : 1 ≤ 3. Proof. now apply nonneg_plus_le_compat_r, le_0_2. Qed. Lemma le_1_4 : 1 ≤ 4. Proof. now apply nonneg_plus_le_compat_r, le_0_3. Qed. Lemma le_2_3 : 2 ≤ 3. Proof. apply (order_preserving (1+)), le_1_2. Qed. Lemma le_2_4 : 2 ≤ 4. Proof. apply (order_preserving (1+)), le_1_3. Qed. Lemma le_3_4 : 3 ≤ 4. Proof. apply (order_preserving (1+)), le_2_3. Qed. Lemma ge_1_mult_compat x y : 1 ≤ x → 1 ≤ y → 1 ≤ x * y. Proof. intros. apply ge_1_mult_le_compat_r; trivial. transitivity 1. now apply le_0_1. easy. Qed. Lemma gt_1_ge_1_mult_compat x y : 1 < x → 1 ≤ y → 1 < x * y. Proof. intros. apply lt_le_trans with x; trivial. apply ge_1_mult_le_compat_r; try easy. transitivity 1. now apply le_0_1. now apply lt_le. Qed. Lemma ge_1_gt_1_mult_compat x y : 1 ≤ x → 1 < y → 1 < x * y. Proof. intros. rewrite commutativity. now apply gt_1_ge_1_mult_compat. Qed. Context `{PropHolds (1 ≶ 0)}. Lemma not_le_1_0 : ¬1 ≤ 0. Proof. now apply lt_not_le_flip, lt_0_1. Qed. Lemma not_le_2_0 : ¬2 ≤ 0. Proof. now apply lt_not_le_flip, lt_0_2. Qed. End full_pseudo_semiring_order. (* Due to bug #2528 *) #[global] Hint Extern 7 (PropHolds (0 ≤ 1)) => eapply @le_0_1 : typeclass_instances. #[global] Hint Extern 7 (PropHolds (0 ≤ 2)) => eapply @le_0_2 : typeclass_instances. #[global] Hint Extern 7 (PropHolds (0 ≤ 3)) => eapply @le_0_3 : typeclass_instances. #[global] Hint Extern 7 (PropHolds (0 ≤ 4)) => eapply @le_0_4 : typeclass_instances. Section dec_semiring_order. (* Maybe these assumptions can be weakened? *) Context `{SemiRingOrder R} `{Apart R} `{!TrivialApart R} `{!NoZeroDivisors R} `{!TotalRelation (≤)} `{∀ x y, Decision (x = y)}. Context `{Rlt : Lt R} (lt_correct : ∀ x y, x < y ↔ x ≤ y ∧ x ≠ y). Instance: FullPseudoOrder (_ : Le R) (_ : Lt R) := dec_full_pseudo_order lt_correct. Instance: StrongSetoid R := pseudo_order_setoid. Instance dec_pseudo_srorder: PseudoSemiRingOrder (<). Proof. split; try apply _. intros x y E. now apply srorder_partial_minus, not_lt_le_flip. intros z. repeat (split; try apply _). intros x y. rewrite !lt_correct. intros [E2a E2b]. split. now apply (order_preserving (z+)). intros E3. apply E2b. now apply (left_cancellation (+) z). now apply (strong_setoids.dec_strong_binary_morphism (.*.)). intros x y. rewrite !lt_correct. intros [E1a E1b] [E2a E2b]. split. now apply nonneg_mult_compat. apply not_symmetry. now apply mult_ne_0; apply not_symmetry. Qed. Instance dec_full_pseudo_srorder: FullPseudoSemiRingOrder (≤) (<). Proof. split; try apply _. now apply le_iff_not_lt_flip. Qed. End dec_semiring_order. Section another_semiring. Context `{SemiRingOrder R1}. Lemma projected_srorder `{SemiRing R2} `{R2le : Le R2} (f : R2 → R1) `{!SemiRing_Morphism f} `{!Injective f} : (∀ x y, x ≤ y ↔ f x ≤ f y) → (∀ x y : R2, x ≤ y → ∃ z, y = x + z) → SemiRingOrder R2le. Proof. intros P. pose proof (projected_partial_order f P). repeat (split; try apply _). assumption. intros. apply P. rewrite 2!preserves_plus. now apply (order_preserving _), P. intros. apply P. apply (order_reflecting (f z +)). rewrite <-2!preserves_plus. now apply P. intros. apply P. rewrite preserves_mult, preserves_0. now apply nonneg_mult_compat; rewrite <-(preserves_0 (f:=f)); apply P. Qed. Context `{SemiRingOrder R2} `{!SemiRing_Morphism (f : R1 → R2)}. (* If a morphism agrees on the positive cone then it is order preserving *) Lemma preserving_preserves_nonneg : (∀ x, 0 ≤ x → 0 ≤ f x) → OrderPreserving f. Proof. intros E. repeat (split; try apply _). intros x y F. destruct (decompose_le F) as [z [Ez1 Ez2]]. apply compose_le with (f z). now apply E. now rewrite Ez2, preserves_plus. Qed. Instance preserves_nonneg `{!OrderPreserving f} x : PropHolds (0 ≤ x) → PropHolds (0 ≤ f x). Proof. intros. rewrite <-(preserves_0 (f:=f)). now apply (order_preserving f). Qed. Lemma preserves_nonpos `{!OrderPreserving f} x : x ≤ 0 → f x ≤ 0. Proof. intros. rewrite <-(preserves_0 (f:=f)). now apply (order_preserving f). Qed. Lemma preserves_ge_1 `{!OrderPreserving f} x : 1 ≤ x → 1 ≤ f x. Proof. intros. rewrite <-(preserves_1 (f:=f)). now apply (order_preserving f). Qed. Lemma preserves_le_1 `{!OrderPreserving f} x : x ≤ 1 → f x ≤ 1. Proof. intros. rewrite <-(preserves_1 (f:=f)). now apply (order_preserving f). Qed. End another_semiring. Section another_semiring_strict. Context `{StrictSemiRingOrder R1} `{StrictSemiRingOrder R2} `{!SemiRing_Morphism (f : R1 → R2)}. Lemma strictly_preserving_preserves_pos : (∀ x, 0 < x → 0 < f x) → StrictlyOrderPreserving f. Proof. intros E. repeat (split; try apply _). intros x y F. destruct (decompose_lt F) as [z [Ez1 Ez2]]. apply compose_lt with (f z). now apply E. now rewrite Ez2, preserves_plus. Qed. Instance preserves_pos `{!StrictlyOrderPreserving f} x : PropHolds (0 < x) → PropHolds (0 < f x). Proof. intros. rewrite <-(preserves_0 (f:=f)). now apply (strictly_order_preserving f). Qed. Lemma preserves_neg `{!StrictlyOrderPreserving f} x : x < 0 → f x < 0. Proof. intros. rewrite <-(preserves_0 (f:=f)). now apply (strictly_order_preserving f). Qed. Lemma preserves_gt_1 `{!StrictlyOrderPreserving f} x : 1 < x → 1 < f x. Proof. intros. rewrite <-(preserves_1 (f:=f)). now apply (strictly_order_preserving f). Qed. Lemma preserves_lt_1 `{!StrictlyOrderPreserving f} x : x < 1 → f x < 1. Proof. intros. rewrite <-(preserves_1 (f:=f)). now apply (strictly_order_preserving f). Qed. End another_semiring_strict. (* Due to bug #2528 *) #[global] Hint Extern 15 (PropHolds (_ ≤ _ _)) => eapply @preserves_nonneg : typeclass_instances. #[global] Hint Extern 15 (PropHolds (_ < _ _)) => eapply @preserves_pos : typeclass_instances. (* Oddly enough, the above hints do not work for goals of the following shape? *) #[global] Hint Extern 15 (PropHolds (_ ≤ '_)) => eapply @preserves_nonneg : typeclass_instances. #[global] Hint Extern 15 (PropHolds (_ < '_)) => eapply @preserves_pos : typeclass_instances. math-classes-8.19.0/quote/000077500000000000000000000000001460576051100153375ustar00rootroot00000000000000math-classes-8.19.0/quote/classquote.v000066400000000000000000000324061460576051100177160ustar00rootroot00000000000000Require Import Coq.Classes.Morphisms Coq.Program.Program Coq.Unicode.Utf8. (* First, two ways to do quoting in the naive scenario without holes/variables in the expression: *) Module simple. (* An example term language and evaluation: *) Inductive Expr := Plus (a b: Expr) | Mult (a b: Expr) | Zero | One. Fixpoint eval (e: Expr): nat := match e with | Plus a b => eval a + eval b | Mult a b => eval a * eval b | Zero => 0 | One => 1 end. (* First up is the simplest approach I can think of. *) Module approach_A. Class Quote (n: nat) := quote: Expr. Arguments quote _ {Quote}. Section instances. Context n m `{Quote n} `{Quote m}. Global Instance: Quote 0 := Zero. Global Instance: Quote 1 := One. Global Instance: Quote (n + m) := Plus (quote n) (quote m). Global Instance: Quote (n * m) := Mult (quote n) (quote m). End instances. Ltac do_quote := match goal with |- (?a = ?b) => change (eval (quote a) = eval (quote b)) end. Lemma example: (1 + 0 + 1) * (1 + 1) = (1 + 1) + (1 + 1). do_quote. Admitted. End approach_A. (* This works, but there's something unsatisfying about this quotation, because the actual Quote instances are not validated until we get to the [change] tactic, which validates the quotation by requiring convertibility. Next, we show an alternative implementation where the Quote instances are all proved locally correct at their definition: *) Module approach_B. Class Quote (n: nat) := { quote: Expr; eval_quote: n = eval quote }. Arguments quote _ {Quote}. Arguments eval_quote _ {Quote}. Section instances. Context n m `{Quote n} `{Quote m}. Global Program Instance: Quote 0 := { quote := Zero }. Global Program Instance: Quote 1 := { quote := One }. Global Instance: Quote (n + m). Proof. refine {| quote := Plus (quote n) (quote m) |}. simpl. do 2 rewrite <- eval_quote. reflexivity. Qed. Global Instance: Quote (n * m). Proof. refine {| quote := Mult (quote n) (quote m) |}. simpl. do 2 rewrite <- eval_quote. reflexivity. Qed. End instances. Lemma do_quote {n m} `{Quote n} `{Quote m}: eval (quote n) = eval (quote m) → n = m. Proof. intros. rewrite (eval_quote n), (eval_quote m). assumption. Qed. Lemma example: (1 + 0 + 1) * (1 + 1) = (1 + 1) + (1 + 1). apply do_quote. Admitted. End approach_B. End simple. (* So far so good, but the variable-less scenario isn't really interesting. We now rework approach B to include quotation of holes/variables, including recognition of syntactically identical ones. *) Module with_vars. (* Some random utilities: *) Lemma sum_assoc {A B C}: (A+B)+C → A+(B+C). intuition. Defined. Lemma bla {A B C}: (A+B) → A+(B+C). intuition. Defined. Lemma monkey {A B}: False + A → A + B. intuition. Defined. Section obvious. Class Obvious (T: Type) := obvious: T. Context (A B C: Type). Global Instance: Obvious (A → A) := id. Global Instance: Obvious (False → A) := False_rect _. Global Instance: Obvious (A → A + B) := inl. Global Instance: Obvious (A → B + A) := inr. Global Instance obvious_sum_src `{Obvious (A → C)} `{Obvious (B → C)}: Obvious (A+B → C). repeat intro. intuition. Defined. Global Instance obvious_sum_dst_l `{Obvious (A → B)}: Obvious (A → B+C). repeat intro. intuition. Defined. Global Instance obvious_sum_dst_r `{Obvious (A → B)}: Obvious (A → C+B). repeat intro. intuition. Defined. End obvious. (* Again our example term language, this time without plus/one (they're boring), but with Var added: *) Inductive Expr (V: Type) := Mult (a b: Expr V) | Zero | Var (v: V). Arguments Var {V}. Arguments Zero {V}. Arguments Mult {V}. (* Require Import monads canonical_names. Instance: MonadReturn Expr := fun _ => Var. Instance expr_bind: MonadBind Expr := fun A B => fix F (m: Expr A) (f: A → Expr B): Expr B := match m with | Zero => Zero | Mult x y => Mult (F x f) (F y f) | Var v => f v end. Section eqs. Context `{e: Equiv A} `{Equivalence _ e}. Global Instance expr_eq: Equiv (Expr A) := fix F (x y: Expr A) := match x, y with | Var v, Var w => v = w | Mult v w, Mult p q => F v p ∧ F w q | Zero, Zero => True | _, _ => False end. Instance: Reflexive expr_eq. Proof. intro. induction x; simpl; intuition. Qed. Instance: Symmetric expr_eq. Proof. intro. induction x; destruct y; simpl in *; intuition. Qed. Instance: Transitive expr_eq. Admitted. Global Instance expr_equivalence: Equivalence expr_eq. End eqs. Instance: ∀ `{Equiv A}, Proper ((=) ==> (=)) (ret Expr). repeat intro. assumption. Qed. Instance bind_proper: ∀ `{Equiv A} `{Equiv B}, Proper ((=) ==> pointwise_relation A (=) ==> (=)) (@expr_bind A B). Proof. intros A H B H0 x y E. (* induction x. destruct y; intuition. intros f g E'. simpl. red. simpl. split. red in E. simpl in E. apply IHx2. simpl in *. unfold expr_bind. simpl. *) Admitted. Instance: Monad Expr. *) (* The expression type is parameterized over the set of variable indices. Hence, we diverge from Claudio, who uses nat indices for variables, thereby introducing bounds problems and dummy variables and other nastiness. *) (* An expression is only meaningful in the context of a variable assignment: *) Definition Value := nat. Definition Vars V := V → Value. Fixpoint eval {V} (vs: Vars V) (e: Expr V): Value := match e with | Zero => 0 | Mult a b => eval vs a * eval vs b | Var v => vs v end. #[global] Instance eval_proper V: Proper (pointwise_relation _ eq ==> eq ==> eq) (@eval V). Proof. repeat intro. subst. induction y0; simpl. congruence. reflexivity. apply H. Qed. (* Some simple combinators for variable packs: *) Definition novars: Vars False := False_rect _. Definition singlevar (x: Value): Vars unit := fun _ => x. Definition merge {A B} (a: Vars A) (b: Vars B): Vars (A+B) := fun i => match i with inl j => a j | inr j => b j end. (* These last two combinators are the "constructors" of an implicitly defined subset of Gallina terms (representing Claudio's "heaps") for which we implement syntactic lookup with type classes: *) Section Lookup. (* Given a heap and value, Lookup instances give the value's index in the heap: *) Class Lookup {A} (x: Value) (f: Vars A) := { lookup: A; lookup_correct: f lookup = x }. Global Arguments lookup {A} _ _ {Lookup}. Context (x: Value) {A B} (va: Vars A) (vb: Vars B). (* If the heap is a merge of two heaps and we can find the value's index in the left heap, we can access it by indexing the merged heap: *) Global Instance lookup_left `{!Lookup x va}: Lookup x (merge va vb). Proof. refine {| lookup := inl (lookup x va) |}. apply lookup_correct. Defined. (* And vice-versa: *) Global Instance lookup_right `{!Lookup x vb} : Lookup x (merge va vb). Proof. refine {| lookup := inr (lookup x vb) |}. apply lookup_correct. Defined. (* If the heap is just a singlevar, we can easily index it. *) Global Program Instance: Lookup x (singlevar x) := { lookup := tt }. (* Note that we don't have any fallback/default instances at this point. We /will/ introduce such an instance for our Quote class later on, which will add a new variable to the heap if another Quote instance that relies on Lookup into the "current" heap fails. *) End Lookup. (* One useful operation we need before we get to Quote relates to variables and expression evaluation. As its name suggests, map_var maps an expression's variable indices. *) Definition map_var {V W: Type} (f: V → W): Expr V → Expr W := fix F (e: Expr V): Expr W := match e with | Mult a b => Mult (F a) (F b) | Zero => Zero | Var v => Var (f v) end. (* An obvious identity is: *) Lemma eval_map_var {V W} (f: V → W) v e: eval v (map_var f e) = eval (v ∘ f) e. Proof. induction e; simpl; try reflexivity. rewrite IHe1, IHe2. reflexivity. Qed. (* Finally, Quote itself: *) Section Quote. (* In Quote, the idea is that V, l, and n are all "input" variables, while V' and r are "output" variables (in the sense that we will rely on unification to generate them. V and l represent the "current heap", n represents the value we want to quote, and V' and r' represent the heap of newly encountered variables during the quotation. This explains the type of quote: it is an expression that refers either to variables from the old heap, or to newly encountered variables. eval_quote is the usual correctness property, which now merges the two heaps. *) Class Quote {V} (l: Vars V) (n: Value) {V'} (r: Vars V') := { quote: Expr (V + V') ; eval_quote: @eval (V+V') (merge l r) quote = n }. Arguments quote {V l} _ {V' r Quote}. (* Our first instance for Zero is easy. The "novars" in the result type reflects the fact that no new variables are encountered. The correctness proof is easy enough for Program. *) Global Program Instance quote_zero V (v: Vars V): Quote v 0 novars := { quote := Zero }. (* The instance for multiplication is a bit more complex. The first line is just boring variable declarations. The second line is important. "Quote x y z" must be read as "quoting y with existing heap x generates new heap z", so the second line basically just shuffles heaps around. The third line has some ugly map_var's in it because the heap shuffling must be reflected in the variable indices, but apart from that it's just constructing a Mult term with quoted subterms. *) Global Program Instance quote_mult V (v: Vars V) n V' (v': Vars V') m V'' (v'': Vars V'') `{!Quote v n v'} `{!Quote (merge v v') m v''}: Quote v (n * m) (merge v' v'') := { quote := Mult (map_var bla (quote n)) (map_var sum_assoc (quote m)) }. Next Obligation. Proof with auto. destruct Quote0, Quote1. subst. simpl. do 2 rewrite eval_map_var. f_equal; apply eval_proper; auto; intro; intuition. Qed. (* Now follows the instance where we recognize values that are already in the heap. This is expressed by the Lookup requirement, which will only be fulfilled if the Lookup instances defined above can find the value in the heap. The novars in the [Quote v x novars] result reflects that this quotation does not generate new variables. *) Global Program Instance quote_old_var V (v: Vars V) x {i: Lookup x v}: Quote v x novars | 8 := { quote := Var (inl (lookup x v)) }. Next Obligation. Proof. apply lookup_correct. Qed. (* Finally, the instance for new variables. We give this lower priority so that it is only used if Lookup fails. *) Global Program Instance quote_new_var V (v: Vars V) x: Quote v x (singlevar x) | 9 := { quote := Var (inr tt) }. End Quote. (* Note: Explicitly using dynamically configured variable index sets instead of plain lists not only removes the need for an awkward dummy value to cope with out-of-bounds accesses, but also means that we can prove the correctness class fields in Lookup/Quote without having to take the potential for out-of-bounds indexing into account (which would be a nightmare). *) (* When quoting something from scratch we will want to start with an empty heap. To avoid having to mention this, we define quote' and eval_quote': *) Definition quote': ∀ x {V'} {v: Vars V'} {d: Quote novars x v}, Expr _ := @quote _ _. Definition eval_quote': ∀ x {V'} {v: Vars V'} {d: Quote novars x v}, eval (merge novars v) quote = x := @eval_quote _ _ . Arguments quote' _ {V' v d}. Arguments eval_quote' _ {V' v d}. (* Time for some tests! *) Goal ∀ x y (P: Value → Prop), P ((x * y) * (x * 0)). intros. rewrite <- (eval_quote' _). (* turns the goal into P (eval some_variable_pack_composed_from_combinators quote) *) simpl quote. Admitted. (* We can also inspect quotations more directly: *) Section inspect. Variables x y: Value. (* Eval compute in quote' ((x * y) * (x * 0)). *) (* = Mult (Mult (Var (inr (inl (inl ())))) (Var (inr (inl (inr ()))))) (Mult (Var (inr (inl (inl ())))) Zero) : Expr (False + (() + () + (False + False))) *) (* The second occurrence of (Var (inr (inl (inl ())))) means the quoting has successfully noticed that it's the same expression. *) (* The two units in the generated variable index type reflect the fact that the expression contains two variables. *) (* I think adding some additional Quote instances might let us get rid of the False's, but at the moment I see little reason to. *) End inspect. (* If we want to quote an equation between two expressions we should make sure that the both sides refer to the same variable pack, and for that we write a little utility function. It does the same kind of shuffling that the mult Quote instance did. *) Lemma quote_equality {V} {v: Vars V} {V'} {v': Vars V'} (l r: Value) `{!Quote novars l v} `{!Quote v r v'}: let heap := (merge v v') in eval heap (map_var monkey quote) = eval heap quote → l = r. Proof with intuition. destruct Quote0 as [lq []]. destruct Quote1 as [rq []]. intros heap H. subst heap. simpl in H. rewrite <- H, eval_map_var. apply eval_proper... intro... Qed. Goal ∀ x y, x * y = y * x. intros. apply (quote_equality _ _). simpl quote. unfold map_var, monkey, sum_rect. Admitted. End with_vars. math-classes-8.19.0/site_scons/000077500000000000000000000000001460576051100163535ustar00rootroot00000000000000math-classes-8.19.0/site_scons/site_tools/000077500000000000000000000000001460576051100205375ustar00rootroot00000000000000math-classes-8.19.0/site_scons/site_tools/Coq.py000066400000000000000000000014731460576051100216400ustar00rootroot00000000000000# -*- coding: utf-8 -*- import SCons.Defaults, SCons.Tool, SCons.Util, os def add_glob(target, source, env): base, _ = os.path.splitext(str(target[0])) target.append(base + ".glob") return target, source Coq = SCons.Builder.Builder( action = '$COQCMD', suffix = '.vo', src_suffix = '.v', emitter = add_glob) def coqdoc_gen(source, target, env, for_signature): for s in source: base, _ = os.path.splitext(str(s)) env.Depends(target, env.File(base + '.glob')) return [SCons.Defaults.Mkdir(target), 'coqdoc $COQDOCFLAGS -d $TARGET $COQFLAGS $SOURCES'] CoqDoc = SCons.Builder.Builder(generator = coqdoc_gen) def generate(env): env['COQC'] = 'coqc' env['COQCMD'] = '$COQC $COQFLAGS -q $SOURCE' env.Append(BUILDERS = {'Coq': Coq, 'CoqDoc': CoqDoc}) def exists(env): return env.Detect('coqc') math-classes-8.19.0/theory/000077500000000000000000000000001460576051100155145ustar00rootroot00000000000000math-classes-8.19.0/theory/CoqStreams.v000066400000000000000000000133121460576051100177640ustar00rootroot00000000000000(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) (* Stream -> Stream. Definition hd (x:Stream) := match x with | Cons a _ => a end. Definition tl (x:Stream) := match x with | Cons _ s => s end. Fixpoint Str_nth_tl (n:nat) (s:Stream) : Stream := match n with | O => s | S m => Str_nth_tl m (tl s) end. Definition Str_nth (n:nat) (s:Stream) : A := hd (Str_nth_tl n s). Lemma tl_nth_tl : forall (n:nat) (s:Stream), tl (Str_nth_tl n s) = Str_nth_tl n (tl s). Proof. simple induction n; simpl; auto. Qed. Hint Resolve tl_nth_tl: datatypes. Lemma Str_nth_tl_plus : forall (n m:nat) (s:Stream), Str_nth_tl n (Str_nth_tl m s) = Str_nth_tl (n + m) s. simple induction n; simpl; intros; auto with datatypes. rewrite <- H. rewrite tl_nth_tl; trivial with datatypes. Qed. Lemma Str_nth_plus : forall (n m:nat) (s:Stream), Str_nth n (Str_nth_tl m s) = Str_nth (n + m) s. intros; unfold Str_nth; rewrite Str_nth_tl_plus; trivial with datatypes. Qed. (** Extensional Equality between two streams *) CoInductive EqSt (s1 s2: Stream) : Prop := eqst : hd s1 = hd s2 -> EqSt (tl s1) (tl s2) -> EqSt s1 s2. (** A coinduction principle *) Ltac coinduction proof := cofix proof; intros; constructor; [ clear proof | try (apply proof; clear proof) ]. (** Extensional equality is an equivalence relation *) Theorem EqSt_reflex : forall s:Stream, EqSt s s. coinduction EqSt_reflex. reflexivity. Qed. Theorem sym_EqSt : forall s1 s2:Stream, EqSt s1 s2 -> EqSt s2 s1. coinduction Eq_sym. case H; intros; symmetry ; assumption. case H; intros; assumption. Qed. Theorem trans_EqSt : forall s1 s2 s3:Stream, EqSt s1 s2 -> EqSt s2 s3 -> EqSt s1 s3. coinduction Eq_trans. transitivity (hd s2). case H; intros; assumption. case H0; intros; assumption. apply (Eq_trans (tl s1) (tl s2) (tl s3)). case H; trivial with datatypes. case H0; trivial with datatypes. Qed. (** The definition given is equivalent to require the elements at each position to be equal *) Theorem eqst_ntheq : forall (n:nat) (s1 s2:Stream), EqSt s1 s2 -> Str_nth n s1 = Str_nth n s2. unfold Str_nth; simple induction n. intros s1 s2 H; case H; trivial with datatypes. intros m hypind. simpl. intros s1 s2 H. apply hypind. case H; trivial with datatypes. Qed. Theorem ntheq_eqst : forall s1 s2:Stream, (forall n:nat, Str_nth n s1 = Str_nth n s2) -> EqSt s1 s2. coinduction Equiv2. apply (H 0). intros n; apply (H (S n)). Qed. Section Stream_Properties. Variable P : Stream -> Prop. (*i Inductive Exists : Stream -> Prop := | Here : forall x:Stream, P x -> Exists x | Further : forall x:Stream, ~ P x -> Exists (tl x) -> Exists x. i*) Inductive Exists ( x: Stream ) : Prop := | Here : P x -> Exists x | Further : Exists (tl x) -> Exists x. CoInductive ForAll (x: Stream) : Prop := HereAndFurther : P x -> ForAll (tl x) -> ForAll x. Lemma ForAll_Str_nth_tl : forall m x, ForAll x -> ForAll (Str_nth_tl m x). Proof. induction m. tauto. intros x [_ H]. simpl. apply IHm. assumption. Qed. Section Co_Induction_ForAll. Variable Inv : Stream -> Prop. Hypothesis InvThenP : forall x:Stream, Inv x -> P x. Hypothesis InvIsStable : forall x:Stream, Inv x -> Inv (tl x). Theorem ForAll_coind : forall x:Stream, Inv x -> ForAll x. coinduction ForAll_coind; auto. Qed. End Co_Induction_ForAll. End Stream_Properties. End Streams. Section Map. Variables A B : Type. Variable f : A -> B. CoFixpoint map (s:Stream A) : Stream B := Cons (f (hd s)) (map (tl s)). Lemma Str_nth_tl_map : forall n s, Str_nth_tl n (map s)= map (Str_nth_tl n s). Proof. induction n. reflexivity. simpl. intros s. apply IHn. Qed. Lemma Str_nth_map : forall n s, Str_nth n (map s)= f (Str_nth n s). Proof. intros n s. unfold Str_nth. rewrite Str_nth_tl_map. reflexivity. Qed. Lemma ForAll_map : forall (P:Stream B -> Prop) (S:Stream A), ForAll (fun s => P (map s)) S <-> ForAll P (map S). Proof. intros P S. split; generalize S; clear S; cofix ForAll_map; intros S; constructor; destruct H as [H0 H]; firstorder. Qed. Lemma Exists_map : forall (P:Stream B -> Prop) (S:Stream A), Exists (fun s => P (map s)) S -> Exists P (map S). Proof. intros P S H. (induction H;[left|right]); firstorder. Defined. End Map. Section Constant_Stream. Variable A : Type. Variable a : A. CoFixpoint const : Stream A := Cons a const. End Constant_Stream. Section Zip. Variable A B C : Type. Variable f: A -> B -> C. CoFixpoint zipWith (a:Stream A) (b:Stream B) : Stream C := Cons (f (hd a) (hd b)) (zipWith (tl a) (tl b)). Lemma Str_nth_tl_zipWith : forall n (a:Stream A) (b:Stream B), Str_nth_tl n (zipWith a b)= zipWith (Str_nth_tl n a) (Str_nth_tl n b). Proof. induction n. reflexivity. intros [x xs] [y ys]. unfold Str_nth in *. simpl in *. apply IHn. Qed. Lemma Str_nth_zipWith : forall n (a:Stream A) (b:Stream B), Str_nth n (zipWith a b)= f (Str_nth n a) (Str_nth n b). Proof. intros. unfold Str_nth. rewrite Str_nth_tl_zipWith. reflexivity. Qed. End Zip. Unset Implicit Arguments. math-classes-8.19.0/theory/abs.v000066400000000000000000000061441460576051100164550ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.orders.rings. Section contents. Context `{Ring R} `{Apart R} `{!TrivialApart R} `{!FullPseudoSemiRingOrder Rle Rlt} `{∀ x y, Decision (x = y)} `{a : !Abs R}. Add Ring R : (rings.stdlib_ring_theory R). Global Instance abs_proper: Setoid_Morphism abs | 0. Proof. split; try apply _. intros x y E. unfold abs, abs_sig. destruct (a x) as [z1 [Ez1 Fz1]]. destruct (a y) as [z2 [Ez2 Fz2]]. simpl. rewrite <-E in Ez2, Fz2. destruct (total (≤) 0 x). now rewrite Ez1, Ez2. now rewrite Fz1, Fz2. Qed. Lemma abs_nonneg (x : R) : 0 ≤ x → abs x = x. Proof. intros E. unfold abs, abs_sig. destruct (a x) as [z1 [Ez1 Fz1]]. auto. Qed. Lemma abs_nonpos (x : R) : x ≤ 0 → abs x = -x. Proof. intros E. unfold abs, abs_sig. destruct (a x) as [z1 [Ez1 Fz1]]. auto. Qed. Lemma abs_nonneg_plus (x y : R) : 0 ≤ x → 0 ≤ y → abs (x + y) = abs x + abs y. Proof. intros Ex Ey. rewrite ?abs_nonneg; try easy. now apply nonneg_plus_compat. Qed. Lemma abs_nonpos_plus (x y : R) : x ≤ 0 → y ≤ 0 → abs (x + y) = abs x + abs y. Proof. intros Ex Ey. rewrite ?abs_nonpos; trivial. ring. now apply nonpos_plus_compat. Qed. Lemma abs_0 : abs 0 = 0. Proof. now rewrite abs_nonneg. Qed. Lemma abs_mult (x y : R) : abs (x * y) = abs x * abs y. Proof with try ring; auto. destruct (total (≤) 0 x) as [Ex|Ex]; destruct (total (≤) 0 y) as [Ey|Ey]. rewrite ?abs_nonneg... apply nonneg_mult_compat... rewrite (abs_nonneg x), (abs_nonpos y), (abs_nonpos (x * y))... apply nonneg_nonpos_mult... rewrite (abs_nonpos x), (abs_nonneg y), (abs_nonpos (x * y))... apply nonpos_nonneg_mult... rewrite (abs_nonpos x), (abs_nonpos y), (abs_nonneg (x * y))... apply nonpos_mult... Qed. Lemma abs_1 : abs 1 = 1. Proof. rewrite abs_nonneg. reflexivity. apply le_0_1. Qed. Lemma abs_negate (x : R) : abs (-x) = abs x. Proof with trivial. destruct (total (≤) 0 x). rewrite (abs_nonneg x), abs_nonpos... apply rings.negate_involutive. apply flip_nonneg_negate... rewrite (abs_nonneg (-x)), abs_nonpos... reflexivity. apply flip_nonpos_negate... Qed. End contents. #[global] Program Instance default_abs `{Ring R} `{!SemiRingOrder Rle} `{∀ x y, Decision (x ≤ y)} : Abs R | 20 := λ x, if decide_rel (≤) 0 x then x else (-x). Next Obligation. split; intros E; try reflexivity. setoid_replace x with 0 by now apply (antisymmetry (≤)). now rewrite rings.negate_0. Qed. Next Obligation. intuition. Qed. Section order_preserving. Context `{Ring A} `{oA : Le A} `{!SemiRingOrder oA} `{!TotalRelation oA} `{!Abs A} `{Ring B} `{oB : Le B} `{!SemiRingOrder oB} `{!TotalRelation oB} `{!Abs B} {f : A → B} `{!OrderPreserving f} `{!SemiRing_Morphism f}. Lemma preserves_abs x : f (abs x) = abs (f x). Proof. destruct (total (≤) 0 x). rewrite ?abs_nonneg; try easy. now apply semirings.preserves_nonneg. rewrite ?abs_nonpos; try easy. apply rings.preserves_negate. now apply preserves_nonpos. Qed. End order_preserving. math-classes-8.19.0/theory/adjunctions.v000066400000000000000000000170441460576051100202320ustar00rootroot00000000000000(** We prove the equivalence of the two definitions of adjunction. *) Require Import MathClasses.interfaces.abstract_algebra MathClasses.theory.setoids MathClasses.interfaces.functors MathClasses.theory.categories MathClasses.misc.workaround_tactics MathClasses.theory.jections. Require MathClasses.categories.dual. Local Hint Unfold id compose: typeclass_instances. (* todo: move *) Local Existing Instance injective_mor. Local Existing Instance surjective_mor. Lemma equal_because_sole `{Setoid T} (P: T → Prop) x: is_sole P x → forall y z, P y → P z → y = z. Proof. firstorder. Qed. (* todo: move *) Section for_φAdjunction. (* MacLane p79/p80, showing that from an φAdjunction we can make both a ηAdjunction and a ηεAdjunction. *) Context `(φAdjunction). Arguments φ {_ _} _. Instance: forall c d, Bijective (@φ c d) := φ_adjunction_bijective F G. Instance: Functor F F' := φ_adjunction_left_functor F G. Instance: Functor G G' := φ_adjunction_right_functor F G. Instance: Category D := functor_from G. Instance: Category C := functor_to G. (* Waiting for the new proof engine ... *) Lemma φ_adjunction_natural_right_inv `(g: c ⟶ G d) `(h: c' ⟶ c): φ⁻¹ (g ◎ h) = φ⁻¹ g ◎ fmap F h. Proof with try reflexivity; try apply _. intros. apply (injective φ). rewrite surjective_applied... rewrite φ_adjunction_natural_right... rewrite surjective_applied... Qed. Lemma φ_adjunction_natural_left_inv `(g: c ⟶ G d) `(k: d ⟶ d'): φ⁻¹ (fmap G k ◎ g) = k ◎ φ⁻¹ g. Proof with try reflexivity; try apply _. intros. apply (injective φ). rewrite surjective_applied... rewrite φ_adjunction_natural_left... rewrite surjective_applied... Qed. Let η: id ⇛ G ∘ F := λ c, φ (cat_id: F c ⟶ F c). Let ε: F ∘ G ⇛ id := λ d, φ ⁻¹ (cat_id: G d ⟶ G d). Global Instance η_natural: NaturalTransformation η. Proof with try reflexivity; try apply _. intros x' x h. change (φ cat_id ◎ h = fmap G (fmap F h) ◎ φ cat_id). rewrite <- φ_adjunction_natural_left, <- φ_adjunction_natural_right, left_identity, right_identity... Qed. Global Instance: NaturalTransformation ε. Proof with try reflexivity; try apply _. intros d d' f. change ((φ ⁻¹) cat_id ◎ fmap F (fmap G f) = f ◎ (φ ⁻¹) cat_id). rewrite <- φ_adjunction_natural_left_inv, <- φ_adjunction_natural_right_inv, (left_identity (fmap G f)), right_identity... Qed. Lemma φ_in_terms_of_η `(f: F x ⟶ a): φ f = fmap G f ◎ η x. Proof. rewrite <- (right_identity f) at 1. rewrite φ_adjunction_natural_left. reflexivity. apply _. Qed. Lemma φ_in_terms_of_ε `(g: x ⟶ G a): φ⁻¹ g = ε a ◎ fmap F g. Proof. rewrite <- (left_identity g) at 1. apply φ_adjunction_natural_right_inv. Qed. Definition univwit (c : C) (d : D): (c ⟶ G d) → (F c ⟶ d) := φ⁻¹. Instance: ∀ c, UniversalArrow (η c: c ⟶ G (F c)) (univwit c). Proof. unfold univwit. constructor; unfold compose. rewrite <- (φ_in_terms_of_η ((φ ⁻¹) f)). symmetry. apply (surjective_applied _). intros ? E. rewrite E. rewrite <- (φ_in_terms_of_η y). symmetry. apply (bijective_applied _). Qed. Instance φAdjunction_ηAdjunction: ηAdjunction F G η univwit := {}. Instance φAdjunction_ηεAdjunction: ηεAdjunction F G η ε. Proof with try apply _. constructor; try apply _; intro x. rewrite <- @φ_in_terms_of_η. unfold ε. apply (surjective_applied _). rewrite <- @φ_in_terms_of_ε. unfold η. apply (surjective_applied _). Qed. (* On a side note, if we let F and G map between the duals of C and D, the adjunction is reversed: *) Goal @φAdjunction D _ _ _ _ C _ _ _ _ G (dual.fmap_op G) F (dual.fmap_op F) (λ d c, (@φ c d)⁻¹) (λ d c, @φ c d). Proof with try apply _. constructor; intros... pose proof (φ_adjunction_bijective F G)... change (d' ⟶ d) in k. change (d ⟶ G c) in f. change ((φ ⁻¹) (f ◎ k) = (φ ⁻¹) f ◎ fmap F k). apply (injective (@φ d' c)). rewrite (surjective_applied _). rewrite φ_adjunction_natural_right... now rewrite (surjective_applied _). change (c ⟶ c') in h. change (d ⟶ G c) in f. change ((φ ⁻¹) (fmap G h ◎ f) = h ◎ (φ ⁻¹) f). apply (injective (@φ d c')). rewrite (surjective_applied _). rewrite φ_adjunction_natural_left... now rewrite (surjective_applied _). Qed. End for_φAdjunction. Section for_ηAdjunction. Context `(ηAdjunction). Instance: Functor F F' := η_adjunction_left_functor F G. Instance: Functor G G' := η_adjunction_right_functor F G. Instance: Category D := functor_from G. Instance: Category C := functor_to G. Let φ x a (g: F x ⟶ a): (x ⟶ G a) := fmap G g ◎ η x. Instance: ∀ (c: C) (d: D), Inverse (@φ c d) := uniwit. Instance: ∀ x a, Surjective (@φ x a). Proof with try apply _. unfold φ. repeat intro. constructor. intros x0 y E. symmetry. rewrite <- E. apply (η_adjunction_universal F G x). constructor... intros ?? E. rewrite E. reflexivity. Qed. Instance: ∀ x a, Injective (@φ x a). Proof with try reflexivity; try apply _; auto. repeat intro. constructor... unfold φ. repeat intro. apply (equal_because_sole _ _ (η_adjunction_universal F G _ _ (fmap G x0 ◎ η x))); unfold compose... Qed. Instance: ∀ x a, Bijective (@φ x a) := {}. Instance ηAdjunction_φAdjunction: φAdjunction F G φ. Proof with try reflexivity; try apply _. unfold φ. unfold id in *. unfold compose in η. constructor... repeat intro. unfold compose. rewrite (associativity (fmap G k))... rewrite preserves_comp... repeat intro. unfold compose. rewrite preserves_comp... rewrite <- associativity. pose proof (η_adjunction_natural F G c' c h) as P. change (η c ◎ h = fmap G (fmap F h) ◎ η c') in P. rewrite <- P. rewrite (associativity (fmap G f))... Qed. End for_ηAdjunction. Section for_ηεAdjunction. Context `(ηεAdjunction). Instance: Functor F F' := ηε_adjunction_left_functor F G. Instance: Functor G G' := ηε_adjunction_right_functor F G. Instance: Category D := functor_from G. Instance: Category C := functor_to G. Instance: NaturalTransformation η := ηε_adjunction_η_natural F G. Instance: NaturalTransformation ε := ηε_adjunction_ε_natural F G. Let φ c d (f: F c ⟶ d): (c ⟶ G d) := fmap G f ◎ η c. Instance uniwit c d: Inverse (φ c d) := λ f, ε d ◎ fmap F f. Instance ηεAdjunction_ηAdjunction: ηAdjunction F G η uniwit. Proof with try apply _. constructor... unfold uniwit. constructor; unfold compose. rewrite preserves_comp... pose proof (ηε_adjunction_η_natural F G c (G d) f) as P. change (η (G d) ◎ f = fmap G (fmap F f) ◎ η c) in P. rewrite <- associativity. rewrite <- P. rewrite associativity. pose proof (ηε_adjunction_identity_at_G F G d) as Q. simpl in Q. rewrite Q. symmetry. apply left_identity. intros y E. rewrite E. clear E f. rewrite preserves_comp... rewrite associativity. pose proof (ηε_adjunction_ε_natural F G (F c) d y) as P. change (ε d ◎ fmap F (fmap G y) = y ◎ ε (F c)) in P. rewrite P. rewrite <- associativity. pose proof (ηε_adjunction_identity_at_F F G c) as Q. simpl in Q. rewrite Q. symmetry. apply right_identity. Qed. Instance ηεAdjunction_φAdjunction: φAdjunction F G φ. Proof. apply ηAdjunction_φAdjunction, _. Qed. End for_ηεAdjunction. math-classes-8.19.0/theory/categories.v000066400000000000000000000234311460576051100200330ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.theory.setoids MathClasses.interfaces.functors MathClasses.theory.jections. Notation "x ⇛ y" := (∀ a, x a ⟶ y a) (at level 90, right associativity) : mc_scope. (* Transformations (polymorphic arrows). Couldn't find an "arrow with dot over it" unicode character. *) (* Natural transformations: *) Definition id_nat_trans `{Arrows D} `{!CatId D} `(F: C → D): F ⇛ F := λ _, cat_id. Section NaturalTransformation. Context `{Category C} `{Category D} `{!Functor (F: C → D) Fa} `{!Functor (G: C → D) Ga}. Class NaturalTransformation (η: F ⇛ G): Prop := natural: ∀ `(f: x ⟶ y), η y ◎ fmap F f = fmap G f ◎ η x. End NaturalTransformation. Section UniversalArrow. Context `{Category D} `{Category C} `{!Functor (F: D → C) Fa}. Class UniversalArrow `(u: c ⟶ F r) (wit: ∀ `(f: c ⟶ F d), r ⟶ d): Prop := universal_arrow: ∀ (d: D) (f: c ⟶ F d), is_sole ((f =) ∘ (◎ u) ∘ fmap F) (wit f). (* Todo: Consider an operational type class for wit. *) End UniversalArrow. Section adjunction. (* Three definitions of adjunctions. *) Context `{Category C} `{Category D} F `{!Functor (F: C → D) F'} G `{!Functor (G: D → C) G'}. (* 1. The definition based on φ (MacLane p79): *) Class φAdjunction (φ: ∀ `(F c ⟶ d), (c ⟶ G d)) `{∀ c d, Inverse (@φ c d)}: Prop := { φ_adjunction_left_functor: Functor F _ ; φ_adjunction_right_functor: Functor G _ ; φ_adjunction_bijective: ∀ c d, Bijective (@φ c d) ; φ_adjunction_natural_left `(f: F c ⟶ d) `(k: d ⟶ d'): φ (k ◎ f) = fmap G k ◎ φ f ; φ_adjunction_natural_right `(f: F c ⟶ d) `(h: c' ⟶ c): φ (f ◎ fmap F h) = φ f ◎ h }. (* 2. The definition based on a universal η (MacLane p81 Theorem 2 (i)): *) Class ηAdjunction (η: id ⇛ G ∘ F) (uniwit: ∀ `(c ⟶ G d), F c ⟶ d): Prop := { η_adjunction_left_functor: Functor F _ ; η_adjunction_right_functor: Functor G _ ; η_adjunction_natural: NaturalTransformation η ; η_adjunction_universal: ∀ c: C, UniversalArrow (η c: c ⟶ G (F c)) (@uniwit c) }. (* We could symmetically define εAdjunction based on universal ε, as in MacLane p81 Theorem 2 (iii), but that would be boring. *) (* 3. The definition based on η and ε being inverses (MacLane p81 Theorem 2 (v)): *) Class ηεAdjunction (η: id ⇛ G ∘ F) (ε: F ∘ G ⇛ id): Prop := { ηε_adjunction_left_functor: Functor F _ ; ηε_adjunction_right_functor: Functor G _ ; ηε_adjunction_η_natural: NaturalTransformation η ; ηε_adjunction_ε_natural: NaturalTransformation ε ; ηε_adjunction_identity_at_G: ∀ a, fmap G (ε a) ◎ η (G a) = id_nat_trans G a ; ηε_adjunction_identity_at_F: ∀ a, ε (F a) ◎ fmap F (η a) = id_nat_trans F a }. (* We currently don't define adjunctions as characterized in MacLane p81 Theorem 2 (ii) and (iv). *) End adjunction. Section contents. Context `{Category X}. Class Mono `(a: x ⟶ y): Prop := mono: ∀ z (f g: z ⟶ x), a ◎ f = a ◎ g → f = g. Section isomorphy. Definition iso_arrows {x y: X} (a: x ⟶ y) (b: y ⟶ x): Prop := a ◎ b = cat_id ∧ b ◎ a = cat_id. (* todo: product *) Global Instance: HeteroSymmetric (@iso_arrows). Proof. unfold iso_arrows. repeat intro. intuition. Qed. Definition is_iso {x y: X} (a: x ⟶ y): Prop := ex (iso_arrows a). Definition isos_unique (x y: X) (a: x ⟶ y) (b b': y ⟶ x): iso_arrows a b → iso_arrows a b' → b = b'. Proof. intros [P Q] [R S]. rewrite <- left_identity. rewrite <- S, <-associativity, P. apply right_identity. Qed. Definition iso: Equiv X := λ x y, ex (uncurry (@iso_arrows x y)). Definition isoT: X → X → Type := λ x y, sig (uncurry (@iso_arrows x y)). Program Instance: Reflexive iso := λ x, ex_intro _ (cat_id, cat_id) _. Next Obligation. split; apply left_identity. Qed. Instance: Symmetric iso. Proof. intros ? ? [[f f'] ?]. exists (f', f). unfold uncurry. apply (hetero_symmetric). assumption. Qed. Instance: Transitive iso. Proof with assumption. intros ? ? ? [[f f'] [U V]] [[g g'] [W Z]]. exists (g ◎ f, f' ◎ g'). split; simpl in *. rewrite <- associativity, (associativity f f' g'), U, left_identity... rewrite <- associativity, (associativity g' g f), Z, left_identity... Qed. Global Instance iso_equivalence: Equivalence iso := {}. Global Instance iso_setoid: @Setoid X iso := {}. Lemma arrows_between_isomorphic_objects (a b c d: X) (ab: a ⟶ b) (ba: b ⟶ a) (cd: c ⟶ d) (dc: d ⟶ c) (ac: a ⟶ c) (bd: b ⟶ d): iso_arrows ab ba → iso_arrows cd dc → ac ◎ ba = dc ◎ bd → bd ◎ ab = cd ◎ ac. Proof. (* shows that you only need one half of the diagram to commute for the other half to commute as well*) intros [H1 H4] [H2 H5] H3. rewrite <- (left_identity (bd ◎ ab)). rewrite <- H2. rewrite <- associativity. rewrite (associativity dc bd ab). rewrite <- H3. rewrite <- associativity. rewrite H4. rewrite right_identity. reflexivity. Qed. Program Definition refl_arrows (x: X): isoT x x := (cat_id, cat_id). Next Obligation. split; apply left_identity. Qed. End isomorphy. Section initiality. Class InitialArrow (x: X): Type := initial_arrow: ∀ y, x ⟶ y. Class Initial (x: X) `{InitialArrow x}: Prop := initial_arrow_unique: ∀ y f', initial_arrow y = f'. Definition initial (x: X): Type := ∀ y: X, sig (λ a: x ⟶ y, ∀ a': x ⟶ y, a = a'). Lemma initials_unique' (x x': X) `{Initial x} `{Initial x'}: iso_arrows (initial_arrow x': x ⟶ x') (initial_arrow x). Proof with reflexivity. split. rewrite <- (H4 _ cat_id), <- H4... rewrite <- (H2 _ cat_id), <- H2... Qed. Program Lemma initials_unique (x x': X) (a: initial x) (b: initial x'): iso_arrows (a x') (b x). Proof. split. destruct (b x') as [? e1]. rewrite <- e1. apply e1. destruct (a x) as [? e0]. rewrite <- e0. apply e0. Qed. End initiality. Section products. Context {I: Type} (component: I → X). Section def. Context (product: X). Class ElimProduct: Type := tuple_proj: ∀ i, product ⟶ component i. Class IntroProduct: Type := make_tuple: ∀ x, (∀ i, x ⟶ component i) → (x ⟶ product). Class Product `{ElimProduct} `{IntroProduct}: Prop := product_factors: ∀ c ccomp, is_sole (λ h', ∀ i, ccomp i = tuple_proj i ◎ h') (make_tuple c ccomp). Lemma tuple_round_trip `{Product} (x: X) (h: ∀ i, x ⟶ component i) i: tuple_proj i ◎ make_tuple x h = h i. Proof. symmetry. apply product_factors. Qed. End def. Lemma products_unique `{Product c} `{Product c'}: iso_arrows (make_tuple c c' (tuple_proj c')) (make_tuple c' c (tuple_proj c)). Proof with intuition. unfold iso_arrows. revert c H1 H2 H3 c' H4 H5 H6. cut (∀ `{Product x} `{Product y}, make_tuple x y (tuple_proj y) ◎ make_tuple y x (tuple_proj x) = cat_id)... pose proof (proj2 (product_factors x x (tuple_proj x))) as Q. rewrite (Q cat_id)... rewrite Q... rewrite associativity. repeat rewrite tuple_round_trip... rewrite right_identity... Qed. End products. Class Producer: Type := product I: (I → X) → X. Class HasProducts `{Producer} `{∀ I c, ElimProduct c (product I c)} `{∀ I c, IntroProduct c (product I c)}: Prop := makes_products: ∀ I (c: I → X), Product c (product I c). (* Definition binary_product `{Producer} (x y: X): Product (λ b: bool => if b then x else y) := produce _. Definition empty_product `{Producer}: Product (λ f: False => match f with end) := produce _. *) (* Section freedom. Context `{Category B} (forget: X → B) `{!Functor forget forget_arr} (S: B). Section candidate. Context {x} (inject: S ⟶ forget x). Definition proves_free (factor: ∀ x', (S ⟶ forget x') → (x ⟶ x')): Prop := ∀ x' (inject': S ⟶ forget x'), is_sole ((inject' =) ∘ (◎ inject) ∘ fmap forget) (factor _ inject'). Definition free: Prop := ex proves_free. End candidate. Lemma frees_unique (x x': X) (inject: S ⟶ forget x) (inject': S ⟶ forget x') (factor: ∀ z, (S ⟶ forget z) → (x ⟶ z)) (factor': ∀ z, (S ⟶ forget z) → (x' ⟶ z)): proves_free inject factor → proves_free inject' factor' → iso_arrows (factor _ inject') (factor' _ inject). Proof with auto; try reflexivity; try apply _. intros P Q. pose proof (proj1 (P _ inject')) as E. pose proof (proj2 (P _ inject)) as R. pose proof (proj1 (Q _ inject)) as E'. pose proof (proj2 (Q _ inject')) as R'. clear P Q. unfold compose in *. split. rewrite (R' cat_id)... apply (R' (factor _ inject' ◎ factor' _ inject)). rewrite preserves_comp... rewrite <- associativity, <- E'... rewrite preserves_id, left_identity... rewrite (R cat_id)... apply (R (factor' _ inject ◎ factor _ inject')). rewrite preserves_comp... rewrite <- associativity, <- E... rewrite preserves_id, left_identity... Qed. End freedom. *) End contents. (* Lemma freedom_as_adjunction `{Category Base} `{Category Extra} `{!Functor (forget: Extra → Base) forget_arr} `{!Functor (freeF: Base → Extra) free_arr} (eta: id ⇛ forget ∘ freeF) (phi: ∀ x y, (x ⟶ forget y) → (freeF x ⟶ y)) `{!ηAdjunction freeF forget eta phi}: ∀ b, proves_free forget b (eta b) (phi b). Proof. exact (alt_adjunction_η_universal _ _ _ _). Qed. *) Arguments Producer : clear implicits. Arguments HasProducts _ {Arrows Eq CatComp H1 H2 H3} : rename. Arguments product {X Producer I} _. math-classes-8.19.0/theory/cut_minus.v000066400000000000000000000235171460576051100177210ustar00rootroot00000000000000Require MathClasses.orders.semirings. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.additional_operations MathClasses.interfaces.orders MathClasses.orders.minmax. (* * Properties of Cut off Minus *) Section cut_minus_properties. Context `{FullPseudoSemiRingOrder R} `{!TrivialApart R} `{∀ x y, Decision (x = y)} `{!CutMinusSpec R cm}. Local Existing Instance pseudo_srorder_semiring. Add Ring SR: (rings.stdlib_semiring_theory R). Let local_le_flip := @orders.le_flip R _ _. Hint Resolve local_le_flip. Global Instance cut_minus_proper: Proper ((=) ==> (=) ==> (=)) cut_minus | 1. Proof. intros x₁ x₂ E y₁ y₂ F. destruct (total (≤) x₂ y₂). rewrite cut_minus_0, cut_minus_0; try easy. now rewrite E, F. apply (right_cancellation (+) y₂). rewrite cut_minus_le by easy. rewrite <-E, <-F. apply cut_minus_le. now rewrite E, F. Qed. Global Instance cut_minus_mor_1: ∀ x : R, Setoid_Morphism (x ∸) | 0. Proof. split; try apply _. Qed. Global Instance cut_minus_mor_2: ∀ x : R, Setoid_Morphism (∸ x) | 0. Proof. split; try apply _. solve_proper. Qed. Let local_cut_minus_0 := (cut_minus_0). Let local_cut_minus_le := (cut_minus_le). Hint Resolve local_cut_minus_0. Hint Resolve local_cut_minus_le. Lemma cut_minus_diag x : x ∸ x = 0. Proof. now apply cut_minus_0. Qed. Lemma cut_minus_nonneg_0_r x : 0 ≤ x → x ∸ 0 = x. Proof. intros E. rewrite <-(rings.plus_0_r (x ∸ 0)). auto. Qed. Lemma cut_minus_0_l x : 0 ≤ x → 0 ∸ x = 0. Proof. auto. Qed. Lemma cut_minus_nonpos_0_r x : x ≤ 0 → x ∸ 0 = 0. Proof. auto. Qed. Lemma cut_minus_nonneg x y : 0 ≤ x ∸ y. Proof. destruct (total (≤) x y) as [E|E]. apply orders.eq_le. symmetry. now auto. apply (order_reflecting (+ y)). now rewrite cut_minus_le; ring_simplify. Qed. Lemma cut_minus_le_r x y : y ≤ x → x ∸ y + y = x. Proof cut_minus_le x y. Lemma cut_minus_le_l x y : y ≤ x → y + (x ∸ y) = x. Proof. rewrite commutativity. now apply cut_minus_le. Qed. Lemma cut_minus_le_trans x y z : y ≤ x → z ≤ y → (x ∸ y) + (y ∸ z) = x ∸ z. Proof. intros. apply (right_cancellation (+) z). rewrite <-associativity, !cut_minus_le; try easy. now transitivity y. Qed. Hint Resolve cut_minus_le_trans. (* We need y₁ ≤ x₁ ∧ y₂ ≤ x₂, e.g. (5 ∸ 6) + (5 ∸0) = 0 + 5 = 5, whereas (10 ∸ 6) = 4 *) (* This example illustrates that y₁ + y₂ ≤ x₁ + x₂ does not work either. *) Lemma cut_minus_plus_distr x₁ x₂ y₁ y₂ : y₁ ≤ x₁ → y₂ ≤ x₂ → (x₁ ∸ y₁) + (x₂ ∸ y₂) = (x₁ + x₂) ∸ (y₁ + y₂). Proof. intros E F. apply (right_cancellation (+) (y₁ + y₂)). rewrite cut_minus_le. setoid_replace (x₁ ∸ y₁ + (x₂ ∸ y₂) + (y₁ + y₂)) with (((x₁ ∸ y₁) + y₁) + ((x₂ ∸ y₂) + y₂)) by ring. now rewrite !cut_minus_le. now apply semirings.plus_le_compat. Qed. (* We need 0 ≤ x, e.g. (-1) * (2 ∸ 1) = -1, whereas (-2) ∸ (-1) = 0 *) Lemma cut_minus_mult_distr_l x y z : 0 ≤ x → x * (y ∸ z) = x * y ∸ x * z. Proof. intros E. destruct (total (≤) y z). rewrite !cut_minus_0; trivial. ring. now apply (maps.order_preserving_nonneg (.*.) x). apply (right_cancellation (+) (x * z)). rewrite <-distribute_l, !cut_minus_le; try easy. now apply (maps.order_preserving_nonneg (.*.) x). Qed. Lemma cut_minus_mult_distr_r x y z : 0 ≤ x → (y ∸ z) * x = y * x ∸ z * x. Proof. intros E. rewrite 3!(commutativity _ x). now apply cut_minus_mult_distr_l. Qed. Lemma cut_minus_plus_rev_l x y z : y ∸ z = (x + y) ∸ (x + z). Proof. destruct (total (≤) y z). rewrite !cut_minus_0; intuition. apply (right_cancellation (+) (x + z)). transitivity ((y ∸ z + z) + x); try ring. rewrite !cut_minus_le; try easy; try ring. now apply (order_preserving (x +)). Qed. Lemma cut_minus_plus_rev_r x y z : y ∸ z = (y + x) ∸ (z + x). Proof. rewrite !(commutativity _ x). now apply cut_minus_plus_rev_l. Qed. (* We need 0 ≤ z, e.g. 2 ∸ (5 - 5) = 2, whereas (2 ∸ 5) ∸ (-5) = 0 ∸ (-5) = 5 *) Lemma cut_minus_plus_r x y z : 0 ≤ z → x ∸ (y + z) = (x ∸ y) ∸ z. Proof. intros E. case (total (≤) x y); intros Exy. rewrite (cut_minus_0 x y) by easy. rewrite cut_minus_0_l, cut_minus_0; try easy. now apply semirings.plus_le_compat_r. rewrite (cut_minus_plus_rev_r y (x ∸ y) z). now rewrite cut_minus_le, commutativity. Qed. Lemma cut_minus_plus_l x y z : 0 ≤ y → x ∸ (y + z) = (x ∸ z) ∸ y. Proof. rewrite commutativity. now apply cut_minus_plus_r. Qed. Lemma cut_minus_plus_toggle1 x y z : x ≤ y → z ≤ y → (y ∸ x) + (x ∸ z) = (y ∸ z) + (z ∸ x). Proof. intros. destruct (total (≤) x z). rewrite (cut_minus_0 x z), cut_minus_le_trans by easy. ring. rewrite (cut_minus_0 z x), cut_minus_le_trans by easy. ring. Qed. Lemma cut_minus_plus_toggle2 x y z : y ≤ x → y ≤ z → (x ∸ z) + (z ∸ y) = (z ∸ x) + (x ∸ y). Proof. intros. destruct (total (≤) x z). rewrite (cut_minus_0 x z), cut_minus_le_trans by easy. ring. rewrite (cut_minus_0 z x) by easy. ring_simplify. now auto. Qed. Lemma cut_minus_plus_toggle3 x₁ x₂ y₁ y₂ : x₁ ≤ y₁ → y₂ ≤ x₂ → (y₁ ∸ x₁) + ((x₁ + x₂) ∸ (y₁ + y₂)) = (x₂ ∸ y₂) + ((y₁ + y₂) ∸ (x₁ + x₂)). Proof. intros. destruct (total (≤) (x₁ + x₂) (y₁ + y₂)). rewrite (cut_minus_0 (x₁ + x₂) (y₁ + y₂)) by easy. rewrite cut_minus_plus_distr by easy. setoid_replace (x₂ + (y₁ + y₂)) with (y₁ + (x₂ + y₂)) by ring. setoid_replace (y₂ + (x₁ + x₂)) with (x₁ + (x₂ + y₂)) by ring. rewrite <-cut_minus_plus_rev_r. ring. rewrite (cut_minus_0 (y₁ + y₂) (x₁ + x₂)) by easy. rewrite cut_minus_plus_distr by easy. setoid_replace (y₁ + (x₁ + x₂)) with (x₂ + (x₁ + y₁)) by ring. setoid_replace (x₁ + (y₁ + y₂)) with (y₂ + (x₁ + y₁)) by ring. rewrite <-cut_minus_plus_rev_r. ring. Qed. Lemma cut_minus_0_plus_toggle x : x + (0 ∸ x) = x ∸ 0. Proof. destruct (total (≤) 0 x). rewrite (cut_minus_0 0 x), (cut_minus_nonneg_0_r x) by easy. ring. rewrite (cut_minus_0 x 0), commutativity; auto. Qed. Lemma cut_minus_0_le x y : x ≤ y → (y ∸ x) + (x ∸ 0) + (0 ∸ y) = (y ∸ 0) + (0 ∸ x). Proof. intros. rewrite <-!cut_minus_0_plus_toggle. apply (right_cancellation (+) x). setoid_replace (y ∸ x + (x + (0 ∸ x)) + (0 ∸ y) + x) with ((y ∸ x + x) + (x + (0 ∸ x)) + (0 ∸ y)) by ring. rewrite (cut_minus_le y x) by easy. ring. Qed. (* * Properties of min and minus *) Section min. Context `{∀ x y : R, Decision (x ≤ y)}. Lemma cut_minus_min1 x y z : x ∸ (y ⊓ z) = x ∸ y + ((x ⊓ y) ∸ z). Proof with eauto; try ring. unfold meet, min, sort. case (decide_rel (≤) x y); case (decide_rel (≤) y z); intros F G; simpl. rewrite (cut_minus_0 x z)... rewrite (cut_minus_0 x y)... rewrite (cut_minus_0 y z)... symmetry... Qed. Lemma cut_minus_min2 x y z : y ∸ z + ((y ⊓ z) ∸ x) = y ∸ x + ((x ⊓ y) ∸ z). Proof. rewrite <-cut_minus_min1. rewrite (commutativity x y), <-cut_minus_min1. now rewrite commutativity. Qed. Lemma cut_minus_min3 x y z : (x + (y ⊓ z)) ∸ ((x + y) ⊓ (x + z)) = ((x + y) ⊓ (x + z)) ∸ (x + (y ⊓ z)). Proof with auto; try reflexivity. destruct (total (≤) y z) as [G1|G1]. rewrite (lattices.meet_l y z), (lattices.meet_l (x + y) (x + z))... rewrite (lattices.meet_r y z), (lattices.meet_r (x + y) (x + z))... Qed. End min. (* The relation to ring minus *) Context `{Negate R} `{!Ring R}. Add Ring R: (rings.stdlib_ring_theory R). Lemma cut_minus_ring_minus (x y : R) : y ≤ x → x ∸ y = x - y. Proof. intros. apply (right_cancellation (+) y). ring_simplify. now auto. Qed. Lemma cut_minus_negate (x : R) : x ≤ 0 → 0 ∸ x = -x. Proof. intros. now rewrite <-(rings.plus_0_l (-x)), cut_minus_ring_minus. Qed. End cut_minus_properties. (* * Default implementation for Rings *) Section cut_minus_default. Context `{Ring R} `{!SemiRingOrder o} `{∀ x y : R, Decision (x ≤ y)}. Global Instance default_cut_minus: CutMinus R | 10 := λ x y, if decide_rel (≤) x y then 0 else x - y. Add Ring R2: (rings.stdlib_ring_theory R). Global Instance: CutMinusSpec R default_cut_minus. Proof. split; unfold cut_minus, default_cut_minus; intros x y ?. case (decide_rel (≤) x y); intros. ring_simplify. now apply (antisymmetry (≤)). ring. now case (decide_rel (≤) x y). Qed. End cut_minus_default. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Opaque default_cut_minus. Set Warnings "+unsupported-attributes". Section order_preserving. Context `{FullPseudoSemiRingOrder A} `{!TrivialApart A} `{!CutMinusSpec A cmA} `{∀ x y : A, Decision (x = y)} `{FullPseudoSemiRingOrder B} `{!TrivialApart B} `{!CutMinusSpec B cmB} `{∀ x y : B, Decision (x = y)} {f : A → B} `{!OrderPreserving f} `{!SemiRing_Morphism f}. Local Existing Instance pseudo_srorder_semiring. Lemma preserves_cut_minus x y : f (x ∸ y) = f x ∸ f y. Proof. destruct (total (≤) x y) as [E|E]. rewrite (cut_minus_0 x y E), (cut_minus_0 (f x) (f y)). apply rings.preserves_0. now apply (order_preserving _). apply (left_cancellation (+) (f y)). rewrite 2!(commutativity (f y)). rewrite <-rings.preserves_plus. rewrite (cut_minus_le x y E), (cut_minus_le (f x) (f y)). reflexivity. now apply (order_preserving _). Qed. End order_preserving. math-classes-8.19.0/theory/dec_fields.v000066400000000000000000000200061460576051100177620ustar00rootroot00000000000000Require Import Coq.setoid_ring.Field Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.vectorspace MathClasses.theory.fields MathClasses.theory.strong_setoids. Require Export MathClasses.theory.rings. Section contents. Context `{DecField F} `{∀ x y: F, Decision (x = y)}. Add Ring F : (stdlib_ring_theory F). Global Instance: ZeroProduct F. Proof with auto. intros x y E. destruct (decide (x = 0)) as [? | Ex]... right. rewrite <-(mult_1_r y), <-(dec_recip_inverse x) by assumption. rewrite associativity, (commutativity y), E. ring. Qed. Global Instance: IntegralDomain F. Proof. split; try apply _. Qed. Lemma dec_recip_1: / 1 = 1. Proof. rewrite <-(rings.mult_1_l (/1)). apply dec_recip_inverse. solve_propholds. Qed. Lemma dec_recip_distr (x y: F): / (x * y) = / x * / y. Proof. destruct (decide (x = 0)) as [Ex|Ex]. rewrite Ex, left_absorb, dec_recip_0. ring. destruct (decide (y = 0)) as [Ey|Ey]. rewrite Ey, right_absorb, dec_recip_0. ring. assert (x * y ≠ 0) as Exy by now apply mult_ne_0. apply (left_cancellation_ne_0 (.*.) (x * y)); trivial. transitivity (x / x * (y / y)). rewrite !dec_recip_inverse by assumption. ring. ring. Qed. Lemma dec_recip_zero x : / x = 0 ↔ x = 0. Proof. split; intros E. apply stable. intros Ex. destruct (is_ne_0 1). rewrite <-(dec_recip_inverse x), E by assumption. ring. rewrite E. now apply dec_recip_0. Qed. Lemma dec_recip_ne_0_iff x : / x ≠ 0 ↔ x ≠ 0. Proof. now split; intros E1 E2; destruct E1; apply dec_recip_zero. Qed. Instance dec_recip_ne_0 x : PropHolds (x ≠ 0) → PropHolds (/x ≠ 0). Proof. intro. now apply dec_recip_ne_0_iff. Qed. Lemma equal_by_one_quotient (x y : F) : x / y = 1 → x = y. Proof. intro Exy. destruct (decide (y = 0)) as [Ey|Ey]. exfalso. apply (is_ne_0 1). rewrite <- Exy, Ey, dec_recip_0. ring. apply (right_cancellation_ne_0 (.*.) (/y)). now apply dec_recip_ne_0. now rewrite dec_recip_inverse. Qed. Global Instance dec_recip_inj: Injective (/). Proof. repeat (split; try apply _). intros x y E. destruct (decide (y = 0)) as [Ey|Ey]. rewrite Ey in *. rewrite dec_recip_0 in E. now apply dec_recip_zero. apply (right_cancellation_ne_0 (.*.) (/y)). now apply dec_recip_ne_0. rewrite dec_recip_inverse by assumption. rewrite <-E, dec_recip_inverse. easy. apply dec_recip_ne_0_iff. rewrite E. now apply dec_recip_ne_0. Qed. Global Instance dec_recip_involutive: Involutive (/). Proof. intros x. destruct (decide (x = 0)) as [Ex|Ex]. now rewrite Ex, !dec_recip_0. apply (right_cancellation_ne_0 (.*.) (/x)). now apply dec_recip_ne_0. rewrite dec_recip_inverse by assumption. rewrite commutativity, dec_recip_inverse. reflexivity. now apply dec_recip_ne_0. Qed. Lemma equal_dec_quotients (a b c d : F) : b ≠ 0 → d ≠ 0 → (a * d = c * b ↔ a / b = c / d). Proof with trivial; try ring. split; intro E. apply (right_cancellation_ne_0 (.*.) b)... apply (right_cancellation_ne_0 (.*.) d)... transitivity (a * d * (b * /b))... transitivity (c * b * (d * /d))... rewrite E, dec_recip_inverse, dec_recip_inverse... transitivity (a * d * 1)... rewrite <-(dec_recip_inverse b)... transitivity (a * / b * d * b)... rewrite E. transitivity (c * (d * / d) * b)... rewrite dec_recip_inverse... Qed. Lemma dec_quotients (a c b d : F) : b ≠ 0 → d ≠ 0 → a / b + c / d = (a * d + c * b) / (b * d). Proof with auto. intros A B. assert (a / b = (a * d) / (b * d)) as E1. apply ->equal_dec_quotients... ring. intros G. destruct (zero_product b d)... assert (c / d = (b * c) / (b * d)) as E2. apply ->equal_dec_quotients... ring. intros G. destruct (zero_product b d)... rewrite E1, E2. ring. Qed. Lemma dec_recip_swap_l x y: x / y = / (/ x * y). Proof. rewrite dec_recip_distr, involutive. ring. Qed. Lemma dec_recip_swap_r x y: / x * y = / (x / y). Proof. rewrite dec_recip_distr, involutive. ring. Qed. Lemma dec_recip_negate x : -(/ x) = / (-x). Proof. destruct (decide (x = 0)) as [Ex|Ex]. now rewrite Ex, negate_0, dec_recip_0, negate_0. apply (left_cancellation_ne_0 (.*.) (-x)). now apply flip_negate_ne_0. rewrite dec_recip_inverse. ring_simplify. now apply dec_recip_inverse. now apply flip_negate_ne_0. Qed. Instance: @VectorSpace F F Ae Aplus Amult Azero Aone Anegate Adec_recip Ae Aplus Azero Anegate Amult. Proof. repeat split; repeat (try apply _). - apply dec_recip_0. - now apply dec_recip_inverse. Qed. End contents. (* Due to bug #2528 *) #[global] Hint Extern 7 (PropHolds (/ _ ≠ 0)) => eapply @dec_recip_ne_0 : typeclass_instances. (* Given a decidable field we can easily construct a constructive field. *) Section is_field. Context `{DecField F} `{Apart F} `{!TrivialApart F} `{∀ x y: F, Decision (x = y)}. Global Program Instance recip_dec_field: Recip F := λ x, /`x. Instance: StrongSetoid F := dec_strong_setoid. Global Instance: Field F. Proof. split; try apply _. now apply (dec_strong_binary_morphism (+)). now apply (dec_strong_binary_morphism (.*.)). split; try apply _. intros ? ? E. unfold recip, recip_dec_field. now apply sm_proper. intros [x Px]. rapply (dec_recip_inverse x). now apply trivial_apart. Qed. Lemma dec_recip_correct (x : F) Px : / x = // x↾Px. Proof with auto. apply (left_cancellation_ne_0 (.*.) x). now apply trivial_apart. now rewrite dec_recip_inverse, reciperse_alt by now apply trivial_apart. Qed. End is_field. Definition stdlib_field_theory F `{DecField F} : Field_theory.field_theory 0 1 (+) (.*.) (λ x y, x - y) (-) (λ x y, x / y) (/) (=). Proof with auto. intros. constructor. apply (theory.rings.stdlib_ring_theory _). apply (is_ne_0 1). reflexivity. intros. rewrite commutativity. now apply dec_recip_inverse. Qed. Section from_stdlib_field_theory. Context `(ftheory : @field_theory F Fzero Fone Fplus Fmult Fminus Fnegate Fdiv Frecip Fe) (rinv_0 : Fe (Frecip Fzero) Fzero) `{!@Setoid F Fe} `{!Proper (Fe ==> Fe ==> Fe) Fplus} `{!Proper (Fe ==> Fe ==> Fe) Fmult} `{!Proper (Fe ==> Fe) Fnegate} `{!Proper (Fe ==> Fe) Frecip}. Add Field F2 : ftheory. Definition from_stdlib_field_theory: @DecField F Fe Fplus Fmult Fzero Fone Fnegate Frecip. Proof with auto. destruct ftheory. repeat (constructor; try assumption); repeat intro ; unfold equiv, mon_unit, sg_op, zero_is_mon_unit, plus_is_sg_op, one_is_mon_unit, mult_is_sg_op, plus, mult, recip, negate; try field... unfold recip, mult. simpl. assert (Fe (Fmult x (Frecip x)) (Fmult (Frecip x) x)) as E by ring. rewrite E... Qed. End from_stdlib_field_theory. Section morphisms. Context `{DecField F} `{∀ x y: F, Decision (x = y)}. Global Instance dec_field_to_domain_inj `{IntegralDomain R} `{!SemiRing_Morphism (f : F → R)} : Injective f. Proof. apply injective_preserves_0. intros x Efx. apply stable. intros Ex. destruct (is_ne_0 (1:R)). rewrite <-(rings.preserves_1 (f:=f)). rewrite <-(dec_recip_inverse x) by assumption. rewrite rings.preserves_mult, Efx. now apply left_absorb. Qed. Lemma preserves_dec_recip `{DecField F2} `{∀ x y: F2, Decision (x = y)} `{!SemiRing_Morphism (f : F → F2)} x : f (/ x) = / f x. Proof. case (decide (x = 0)) as [E | E]. now rewrite E, dec_recip_0, preserves_0, dec_recip_0. intros. apply (left_cancellation_ne_0 (.*.) (f x)). now apply injective_ne_0. rewrite <-preserves_mult, 2!dec_recip_inverse. now apply preserves_1. now apply injective_ne_0. easy. Qed. Lemma dec_recip_to_recip `{Field F2} `{!StrongSemiRing_Morphism (f : F → F2)} x Pfx : f (/ x) = // (f x)↾Pfx. Proof. assert (x ≠ 0). intros Ex. destruct (apart_ne (f x) 0 Pfx). now rewrite Ex, preserves_0. apply (left_cancellation_ne_0 (.*.) (f x)). now apply injective_ne_0. rewrite <-preserves_mult, dec_recip_inverse, reciperse_alt by assumption. now apply preserves_1. Qed. End morphisms. math-classes-8.19.0/theory/fields.v000066400000000000000000000130721460576051100171540ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.theory.strong_setoids. Require Export MathClasses.theory.rings. Section field_properties. Context `{Field F}. Add Ring F : (stdlib_ring_theory F). Lemma reciperse_alt (x : F) Px : x // x↾Px = 1. Proof. now rewrite <-(recip_inverse (x↾Px)). Qed. Lemma recip_proper_alt x y Px Py : x = y → // x↾Px = // y↾Py. Proof. intro. now apply sm_proper. Qed. Lemma recip_irrelevant x Px1 Px2 : // x↾Px1 = // x↾Px2. Proof. now apply recip_proper_alt. Qed. Lemma apart_0_proper {x y} : x ≶ 0 → x = y → y ≶ 0. Proof. intros ? E. now rewrite <-E. Qed. Global Instance: StrongInjective (-). Proof. repeat (split; try apply _); intros x y E. apply (strong_extensionality (+ x + y)). ring_simplify. now symmetry. apply (strong_extensionality (+ -x + -y)). ring_simplify. now symmetry. Qed. Global Instance: StrongInjective (//). Proof. repeat (split; try apply _); intros x y E. apply (strong_extensionality (`x *.)). rewrite recip_inverse, commutativity. apply (strong_extensionality (`y *.)). rewrite associativity, recip_inverse. ring_simplify. now symmetry. apply (strong_extensionality (.* // x)). rewrite recip_inverse, commutativity. apply (strong_extensionality (.* // y)). rewrite <-associativity, recip_inverse. ring_simplify. now symmetry. Qed. Global Instance: ∀ z, StrongLeftCancellation (+) z. Proof. intros z x y E. apply (strong_extensionality (+ -z)). now ring_simplify. Qed. Global Instance: ∀ z, StrongRightCancellation (+) z. Proof. intros. apply (strong_right_cancel_from_left (+)). Qed. Global Instance: ∀ z, PropHolds (z ≶ 0) → StrongLeftCancellation (.*.) z. Proof. intros z Ez x y E. red in Ez. rewrite !(commutativity z). apply (strong_extensionality (.* // z↾(Ez : (≶0) z))). rewrite <-!associativity, !reciperse_alt. now ring_simplify. Qed. Global Instance: ∀ z, PropHolds (z ≶ 0) → StrongRightCancellation (.*.) z. Proof. intros. apply (strong_right_cancel_from_left (.*.)). Qed. Lemma mult_apart_zero_l x y : x * y ≶ 0 → x ≶ 0. Proof. intros. apply (strong_extensionality (.* y)). now rewrite mult_0_l. Qed. Lemma mult_apart_zero_r x y : x * y ≶ 0 → y ≶ 0. Proof. intros. apply (strong_extensionality (x *.)). now rewrite mult_0_r. Qed. Instance mult_apart_zero x y : PropHolds (x ≶ 0) → PropHolds (y ≶ 0) → PropHolds (x * y ≶ 0). Proof. intros Ex Ey. apply (strong_extensionality (.* // y↾(Ey : (≶0) y))). now rewrite <-associativity, reciperse_alt, mult_1_r, mult_0_l. Qed. Instance: NoZeroDivisors F. Proof. intros x [x_nonzero [y [y_nonzero E]]]. rewrite <-tight_apart in y_nonzero. destruct y_nonzero. intro y_apartzero. apply x_nonzero. rewrite <- mult_1_r. rewrite <- (reciperse_alt y y_apartzero). rewrite associativity, E. ring. Qed. Global Instance: IntegralDomain F. Proof. split; try apply _. Qed. Global Instance apart_0_sig_apart_0: ∀ (x : F ₀), PropHolds (`x ≶ 0). Proof. now intros [??]. Qed. Instance recip_apart_zero x : PropHolds (// x ≶ 0). Proof. red. apply mult_apart_zero_r with (`x). rewrite recip_inverse. solve_propholds. Qed. Lemma field_div_0_l x y : x = 0 → x // y = 0. Proof. intros E. rewrite E. apply left_absorb. Qed. Lemma field_div_diag x y : x = `y → x // y = 1. Proof. intros E. rewrite E. apply recip_inverse. Qed. Lemma equal_quotients (a c: F) b d : a * `d = c * `b ↔ a // b = c // d. Proof with try ring. split; intro E. transitivity (1 * (a // b))... rewrite <- (recip_inverse d). transitivity (// d * (a * `d) // b)... rewrite E. transitivity (// d * c * (`b // b))... rewrite recip_inverse... transitivity (a * `d * 1)... rewrite <- (recip_inverse b). transitivity (a // b * `d * `b)... rewrite E. transitivity (c * (`d // d) * `b)... rewrite recip_inverse... Qed. (* todo: should be cleanable *) Lemma recip_distr_alt (x y : F) Px Py Pxy : // (x * y)↾Pxy = // x↾Px * // y↾Py. Proof with try ring. apply (left_cancellation_ne_0 (.*.) (x * y)). now apply apart_ne. transitivity ((x // x↾Px) * (y // y↾Py))... transitivity ((x * y) // (x * y)↾Pxy)... rewrite 3!reciperse_alt... Qed. End field_properties. (* Due to bug #2528 *) #[global] Hint Extern 8 (PropHolds (// _ ≶ 0)) => eapply @recip_apart_zero : typeclass_instances. #[global] Hint Extern 8 (PropHolds (_ * _ ≶ 0)) => eapply @mult_apart_zero : typeclass_instances. Section morphisms. Context `{Field F1} `{Field F2} `{!StrongSemiRing_Morphism (f : F1 → F2)}. Add Ring F1 : (stdlib_ring_theory F1). Lemma strong_injective_preserves_0 : (∀ x, x ≶ 0 → f x ≶ 0) → StrongInjective f. Proof. intros E1. split; try apply _. intros x y E2. apply (strong_extensionality (+ -f y)). rewrite plus_negate_r, <-preserves_minus. apply E1. apply (strong_extensionality (+ y)). now ring_simplify. Qed. (* We have the following for morphisms to non-trivial strong rings as well. However, since we do not have an interface for strong rings, we ignore it. *) Global Instance: StrongInjective f. Proof. apply strong_injective_preserves_0. intros x Ex. apply mult_apart_zero_l with (f (// exist (≶0) x Ex)). rewrite <-rings.preserves_mult. rewrite reciperse_alt. rewrite rings.preserves_1. solve_propholds. Qed. Lemma preserves_recip x Px Pfx : f (// x↾Px) = // (f x)↾Pfx. Proof. apply (left_cancellation_ne_0 (.*.) (f x)). now apply apart_ne. rewrite <-rings.preserves_mult. rewrite !reciperse_alt. now apply preserves_1. Qed. End morphisms. math-classes-8.19.0/theory/finite_sets.v000066400000000000000000000361301460576051100202220ustar00rootroot00000000000000Require Import MathClasses.theory.lattices MathClasses.varieties.monoids MathClasses.implementations.bool MathClasses.implementations.list_finite_set MathClasses.orders.lattices MathClasses.interfaces.abstract_algebra MathClasses.interfaces.finite_sets MathClasses.interfaces.orders. Definition fset_car_setoid A `{FSet A} : Setoid A := setoidmor_a singleton. Section fset_props. Context `{FSet A}. Instance: Setoid A := fset_car_setoid A. Lemma fset_extend_correct_applied `{BoundedJoinSemiLattice B} (f : A → B) `{!Setoid_Morphism f} x : f x = fset_extend f {{ x }}. Proof. now apply fset_extend_correct. Qed. Lemma fset_extend_unique_applied `{Equiv B} `{Join B} `{Bottom B} (f : A → B) `{!Setoid_Morphism f} (h : set_type A → B) `{!BoundedJoinSemiLattice_Morphism (h : set_type A → B)} : (∀ x, f x = h {{ x }}) → ∀ x, h x = fset_extend f x. Proof. intros. apply setoids.ext_equiv_applied, (fset_extend_unique _ _). now apply setoids.ext_equiv_applied_iff. Qed. Let F (x y z : A) := if decide (z = x) then false else true. Instance: ∀ x y, Setoid_Morphism (F x y). Proof. split; try apply _. intros ?? E. unfold F. do 2 case (decide _); try reflexivity; rewrite E; contradiction. Qed. Global Instance: Injective (singleton : A → set_type A). Proof. split; try apply _. intros x y E1. apply stable; intros E2. assert (fset_extend (F x y) {{ x }} ≠ fset_extend (F x y) {{ y }}) as E3. rewrite <-!(fset_extend_correct_applied (F x y)). unfold F. do 2 case (decide _); intuition try discriminate; auto. destruct E3. now rewrite E1. Qed. Lemma fset_singleton_ne_empty x : {{ x }} ≠ ∅. Proof. intros E1. set (g (z : A) := true). assert (Setoid_Morphism g) by (split; try apply _; firstorder). assert (fset_extend g {{ x }} ≠ ⊥) as E2. rewrite <-(fset_extend_correct_applied g). discriminate. destruct E2. now rewrite E1, preserves_bottom. Qed. Lemma fset_join_singletons x : {{ x ; x }} = {{ x }}. Proof. now rewrite (idempotency (⊔) _). Qed. Lemma fset_join_singletons_eq_l x y : {{ x ; y }} = {{ x }} ↔ x = y. Proof. split. intros E1. apply stable; intros E2. assert (fset_extend (F x y) {{ x ; y }} ≠ fset_extend (F x y) {{ x }}) as E3. rewrite preserves_join, <-!(fset_extend_correct_applied (F x y)). unfold F. do 2 (case (decide _)); intuition try discriminate; auto. destruct E3. now rewrite E1. intros E. now rewrite E, fset_join_singletons. Qed. Lemma fset_join_singletons_eq_r x y : {{ x ; y }} = {{ y }} ↔ x = y. Proof. rewrite commutativity, fset_join_singletons_eq_l. intuition. Qed. End fset_props. #[global] Instance fset_map_mor `{FSet A} `{FSet B} (f : A → B) `{!Setoid_Morphism f} : BoundedJoinSemiLattice_Morphism (fset_map (H:=At) (H0:=At0) (SetSingleton0:=Asingle0) f). Proof. apply _. Qed. Lemma fset_map_correct `{FSet A} `{FSet B} (f : A → B) `{!Setoid_Morphism f} : singleton ∘ f = fset_map f ∘ singleton. Proof (fset_extend_correct _). Lemma fset_map_correct_applied `{FSet A} `{FSet B} (f : A → B) `{!Setoid_Morphism f} x : {{ f x }} = fset_map f {{ x }}. Proof. pose proof (fset_car_setoid A). now apply (setoids.ext_equiv_applied (fset_map_correct f)). Qed. Lemma fset_map_unique `{FSet A} `{FSet B} (f : A → B) `{!Setoid_Morphism f} (h : set_type A → set_type B) `{!BoundedJoinSemiLattice_Morphism h} : singleton ∘ f = h ∘ singleton → h = fset_map f. Proof. intros. unfold fset_map. now apply (fset_extend_unique _ _). Qed. Lemma fset_map_id `{FSet A} : fset_map id = id. Proof. pose proof (fset_car_setoid A). symmetry. apply (fset_map_unique id id). now apply setoids.ext_equiv_refl. Qed. Lemma fset_map_id_applied `{FSet A} x : fset_map id x = x. Proof. now apply fset_map_id. Qed. Lemma fset_map_compose `{FSet A} `{FSet B} `{FSet C} (f : B → C) `{!Setoid_Morphism f} (g : A → B) `{!Setoid_Morphism g} : fset_map (f ∘ g) = fset_map f ∘ fset_map g. Proof. pose proof (fset_car_setoid A). symmetry. apply (fset_map_unique (f ∘ g) _). rewrite compose_assoc. rewrite <-(fset_map_correct g). rewrite <-compose_assoc. rewrite (fset_map_correct f). now apply setoids.ext_equiv_refl. Qed. Section fset_map_inverse. Context `{FSet A} `{FSet B} (f : A → B) `{!Inverse f} `{!Bijective f}. Global Instance fset_map_inverse: Inverse (fset_map f) := fset_map (f⁻¹). Instance fset_map_surjective: Surjective (fset_map f). Proof. pose proof (fset_car_setoid A). pose proof (fset_car_setoid B). pose proof (injective_mor f). split; try apply _. unfold inverse, fset_map_inverse. rewrite <-(fset_map_compose _ _). symmetry. apply (fset_map_unique _ _). rewrite (surjective f). now apply setoids.ext_equiv_refl. Qed. End fset_map_inverse. #[global] Instance fset_map_bijective `{FSet A} `{FSet B} (f : A → B) `{!Inverse f} `{!Bijective f} : Bijective (fset_map f). Proof. pose proof (fset_car_setoid A). pose proof (fset_car_setoid B). pose proof (injective_mor f). pose proof (fset_map_surjective f). pose proof (fset_map_surjective (f⁻¹)). repeat (split; try apply _). intros x y E. rewrite <-(jections.surjective_applied (fset_map (f⁻¹)) x). rewrite <-(jections.surjective_applied (fset_map (f⁻¹)) y). now apply sm_proper. Qed. Lemma preserves_in `{FullFSet A} `{FullFSet B} (f : A → B) `{!Inverse f} `{!Bijective f} x X : f x ∈ fset_map f X ↔ x ∈ X. Proof. pose proof (injective_mor f). pose proof (join_sl_mor_preserving (fset_map f)). pose proof (join_sl_mor_reflecting (fset_map f)). rewrite !fset_in_singleton_le. split; intros E. apply (order_reflecting (fset_map f)). now rewrite <-(fset_map_correct_applied f). rewrite (fset_map_correct_applied f). now apply (order_preserving _). Qed. Lemma preserves_notin `{FullFSet A} `{FullFSet B} (f : A → B) `{!Inverse f} `{!Bijective f} x X : f x ∉ fset_map f X ↔ x ∉ X. Proof. split; intros E ?; now apply E, (preserves_in f). Qed. Section full_fset_props. Context `{FullFSet A}. Instance: Setoid A := fset_car_setoid A. Notation to_listset := (fset_map id : set_type A → @set_type _ (listset A)). Notation from_listset := (to_listset⁻¹). Lemma to_listset_preserves_in x X : x ∈ to_listset X ↔ x ∈ X. Proof preserves_in id x X. Lemma fset_induction (P : set_type A → Prop) `{proper : !Proper ((=) ==> iff) P} : P ∅ → (∀ x X, x ∉ X → P X → P ({{ x }} ⊔ X)) → ∀ X, P X. Proof. intros Pempty Padd X. mc_setoid_replace X with (from_listset (to_listset X)) by (symmetry; apply (jections.bijective_applied _)). generalize (to_listset X). apply listset_induction. solve_proper. now rewrite preserves_bottom. intros x l E1 E2. change (P (fset_map id ({{x}} ⊔ l))). rewrite preserves_join, <-(fset_map_correct_applied _ x). apply Padd; auto. intros E3. destruct E1. now apply (preserves_in id x). Qed. Global Instance fset_in_proper : Proper ((=) ==> (=) ==> iff) ((∈): A → set_type A). Proof. intros x y E1 X Y E2. now rewrite !fset_in_singleton_le, E1, E2. Qed. Global Program Instance fset_in_dec_slow: ∀ x X, Decision (x ∈ X) | 50 := λ x X, match decide_rel (∈) x (to_listset X) with left E => left _ | right E => right _ end. Next Obligation. now apply to_listset_preserves_in. Qed. Next Obligation. intros F. destruct E. now apply to_listset_preserves_in. Qed. Lemma fset_notin_empty x : x ∉ ∅. Proof. intro. now apply fset_singleton_ne_empty with x, below_bottom, fset_in_singleton_le. Qed. Lemma fset_in_join X Y x : x ∈ X ⊔ Y ↔ x ∈ X ∨ x ∈ Y. Proof. rewrite <-!to_listset_preserves_in, preserves_join. apply listset_in_join. Qed. Lemma fset_notin_join X Y x : x ∉ X ⊔ Y ↔ x ∉ X ∧ x ∉ Y. Proof. rewrite fset_in_join. tauto. Qed. Lemma fset_in_singleton x : x ∈ {{ x }}. Proof. now rewrite fset_in_singleton_le, join_sl_le_spec, fset_join_singletons. Qed. Lemma fset_in_singleton_eq x y : x ∈ {{ y }} ↔ x = y. Proof. split; intros E. now apply fset_join_singletons_eq_r, join_sl_le_spec, fset_in_singleton_le. rewrite E. apply fset_in_singleton. Qed. Lemma fset_notin_singleton_neq x y : x ∉ {{ y }} ↔ x ≠ y. Proof. now rewrite fset_in_singleton_eq. Qed. Lemma fset_in_add y X x : y ∈ {{ x }} ⊔ X ↔ y = x ∨ y ∈ X. Proof. rewrite fset_in_join. split; intros [?|?]; try tauto. left. now apply fset_in_singleton_eq. left. now apply fset_in_singleton_eq. Qed. Lemma fset_notin_add y X x : y ∉ {{ x }} ⊔ X ↔ y ≠ x ∧ y ∉ X. Proof. rewrite fset_in_add. tauto. Qed. Lemma fset_in_inversion y X x : y ∈ {{ x }} ⊔ X → y = x ∨ y ∈ X. Proof. rewrite fset_in_join. intros [?|?]; try tauto. left. now apply fset_in_singleton_eq. Qed. Lemma fset_le_in X Y : X ≤ Y ↔ ∀ x, x ∈ X → x ∈ Y. Proof. pose proof (join_sl_mor_preserving to_listset). pose proof (join_sl_mor_reflecting to_listset). setoid_rewrite <-to_listset_preserves_in. split; intros E. now apply (order_preserving (to_listset)) in E. now apply (order_reflecting (to_listset)). Qed. Lemma fset_eq_in X Y : X = Y ↔ ∀ x, x ∈ X ↔ x ∈ Y. Proof. setoid_rewrite <-to_listset_preserves_in. split. intros E. change (to_listset X = to_listset Y). now apply sm_proper. intros. now apply (injective (to_listset)). Qed. Instance: Proper ((=) ==> (=) ==> (=)) (⊓). Proof. intros ?? E1 ?? E2. apply fset_eq_in. intros. now rewrite !fset_in_meet, E1, E2. Qed. Instance: Associative (⊓). Proof. repeat intro. apply fset_eq_in. intros. rewrite !fset_in_meet. tauto. Qed. Instance: Commutative (⊓). Proof. repeat intro. apply fset_eq_in. intros. rewrite !fset_in_meet. tauto. Qed. Instance: BinaryIdempotent (⊓). Proof. repeat intro. apply fset_eq_in. intros. rewrite !fset_in_meet. tauto. Qed. Instance: MeetSemiLattice (set_type A). Proof. repeat (split; try apply _). Qed. Global Instance: DistributiveLattice (set_type A). Proof. repeat (split; try apply _); repeat intro; apply fset_eq_in; intro; repeat (rewrite fset_in_meet || rewrite fset_in_join); tauto. Qed. Global Instance: MeetSemiLatticeOrder (≤). Proof. apply alt_Build_MeetSemiLatticeOrder. intros. rewrite fset_le_in, fset_eq_in. setoid_rewrite fset_in_meet. firstorder trivial. Qed. Lemma fset_meet_singletons x : {{ x }} ⊓ {{ x }} = {{ x }}. Proof. now rewrite (idempotency (⊔) _). Qed. Lemma fset_meet_singletons_eq_l x y : {{ x }} ⊓ {{ y }} = {{ x }} ↔ x = y. Proof. split; intros E. apply fset_in_singleton_eq. rewrite fset_eq_in in E. setoid_rewrite fset_in_meet in E. now destruct (proj2 (E x) (fset_in_singleton _)). now rewrite E, fset_meet_singletons. Qed. Lemma fset_meet_singletons_eq_r x y : {{ x }} ⊓ {{ y }} = {{ y }} ↔ x = y. Proof. rewrite commutativity, fset_meet_singletons_eq_l. intuition. Qed. Lemma fset_meet_distinct_singletons (x y: A) : x ≠ y → {{ x }} ⊓ {{ y }} = ∅. Proof. intros E1. apply fset_eq_in. intros z. rewrite fset_in_meet. split. intros [E2 E3]. destruct E1. apply fset_in_singleton_eq in E2. apply fset_in_singleton_eq in E3. now rewrite <-E2, <-E3. intro. now destruct (fset_notin_empty z). Qed. Global Instance: Proper ((=) ==> (=) ==> (=)) (∖). Proof. intros ?? E1 ?? E2. apply fset_eq_in. intros. now rewrite !fset_in_difference, E1, E2. Qed. Global Instance fset_difference_empty_r: RightIdentity (∖) ∅. Proof. intro. apply fset_eq_in. intro. rewrite fset_in_difference. split; intuition. edestruct fset_notin_empty; eassumption. Qed. Global Instance fset_difference_empty_l: LeftAbsorb (∖) ∅. Proof. intro. apply fset_eq_in. intro. rewrite fset_in_difference. split; intuition. edestruct fset_notin_empty; eassumption. Qed. Global Instance diff_meet_distr_r: RightDistribute (∖) (⊓). Proof. intros X Y Z. apply fset_eq_in. intro. repeat (rewrite fset_in_meet || rewrite fset_in_difference). intuition. Qed. Global Instance diff_join_distr_r: RightDistribute (∖) (⊔). Proof. intros X Y Z. apply fset_eq_in. intro. repeat (rewrite fset_in_join || rewrite fset_in_difference). intuition. Qed. Lemma diff_meet_join_diff X Y Z : X ∖ (Y ⊓ Z) = X ∖ Y ⊔ X ∖ Z. Proof. apply fset_eq_in. intro. repeat (rewrite fset_in_join || rewrite fset_in_meet || rewrite fset_in_difference). split; try tauto. intros [??]. case (decide (x ∈ Y)); tauto. Qed. Lemma diff_join_diff_meet X Y Z : X ∖ (Y ⊔ Z) = X ∖ Y ⊓ X ∖ Z. Proof. apply fset_eq_in. intro. repeat (rewrite fset_in_join || rewrite fset_in_meet || rewrite fset_in_difference). tauto. Qed. End full_fset_props. Ltac split_sets := repeat (match goal with | E : _ ∈ ∅ |- _ => apply fset_notin_empty in E; destruct E | E : _ ∈ {{ _ }} |- _ => apply fset_in_singleton_eq in E | E : _ ∉ {{ _ }} |- _ => apply fset_notin_singleton_neq in E | E : _ ∈ _ ⊔ _ |- _ => apply fset_in_join in E; destruct E | E : _ ∉ _ ⊔ _ |- _ => apply fset_notin_join in E; destruct E | E : _ ∈ _ ⊓ _ |- _ => apply fset_in_meet in E; destruct E | |- context [_ ∈ _ ⊔ _] => rewrite !fset_in_join end). Section iso_is_fset. Context `{Setoid A} `{At : SetType A} `{BoundedJoinSemiLattice (set_type A)} `{fsetB : FSet B} `{SetSingleton A} `{!Setoid_Morphism (singleton : A → At)} (A_to_B : A → B) `{!Inverse A_to_B} `{!Bijective A_to_B} (At_to_Bt : set_type A → set_type B) `{!Inverse At_to_Bt} `{!Bijective At_to_Bt} `{!BoundedJoinSemiLattice_Morphism At_to_Bt} `{∀ a₁ a₂ : A, Decision (a₁ = a₂)} (singleton_correct : At_to_Bt ∘ singleton = singleton ∘ A_to_B). Instance: Setoid B := fset_car_setoid B. Lemma singleton_correct_alt : At_to_Bt⁻¹ ∘ singleton = singleton ∘ A_to_B⁻¹. Proof. pose proof (injective_mor A_to_B). pose proof (injective_mor At_to_Bt). apply (jections.injective_compose_cancel At_to_Bt). rewrite <-!compose_assoc. rewrite (surjective At_to_Bt), singleton_correct. rewrite compose_assoc, (surjective A_to_B). rewrite compose_id_left, compose_id_right. now apply setoids.ext_equiv_refl. Qed. Instance iso_is_fset_extend: FSetExtend A := λ C _ _ f, fset_extend (f ∘ A_to_B⁻¹) ∘ At_to_Bt. Instance iso_is_fset: FSet A. Proof. pose proof (injective_mor A_to_B). split; unfold fset_extend, iso_is_fset_extend; try apply _. intros C ? ? ? ? f ?. rewrite compose_assoc, singleton_correct, <-compose_assoc. rewrite <-(fset_extend_correct (f ∘ A_to_B ⁻¹)). rewrite compose_assoc, (jections.bijective _), compose_id_right. now apply setoids.ext_equiv_refl. intros C ? ? ? f ? h ? E1. pose proof (bounded_join_slmor_b (f:=h)). rewrite <-(fset_extend_unique (f ∘ A_to_B⁻¹) (h ∘ At_to_Bt⁻¹)). rewrite compose_assoc, (jections.bijective _), compose_id_right. now apply setoids.ext_equiv_refl. rewrite E1, !compose_assoc, singleton_correct_alt. now apply setoids.ext_equiv_refl. Qed. End iso_is_fset. math-classes-8.19.0/theory/forget_algebra.v000066400000000000000000000030461460576051100206510ustar00rootroot00000000000000(* "Forgetting" an algebra's operations (but keeping the setoid equality) is a trivial functor. This functor should nicely compose with the one forgetting variety laws. *) Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.interfaces.functors MathClasses.theory.ua_homomorphisms MathClasses.theory.categories MathClasses.categories.setoids MathClasses.categories.product MathClasses.categories.algebras. Section contents. Variable sign: Signature. Notation TargetObject := (product.Object (λ _: sorts sign, setoids.Object)). Let TargetArrows: Arrows TargetObject := @product.pa _ (λ _: sorts sign, setoids.Object) (λ _, _: Arrows setoids.Object). (* hm, not happy about this *) Definition object (v: algebras.Object sign): TargetObject := λ i, @setoids.object (v i) (algebras.algebra_equiv sign v i) _. Global Program Instance: Fmap object := λ _ _, id. Global Instance forget: Functor object _. Proof. constructor; try apply _. constructor; try apply _. intros x y E i A B F. rewrite F. now apply E. now repeat intro. intros ? ? f ? g i ? ? E. now rewrite E. Qed. Let hint: ∀ x y, Equiv (object x ⟶ object y). intros. apply _. Defined. (* todo: shouldn't be necessary *) Global Instance mono: ∀ (X Y: algebras.Object sign) (a: X ⟶ Y), Mono (fmap object a) → Mono a. Proof. intros ??? H1 ??? H2 ??. assert (fmap object f = fmap object g). apply H1. intros ??? H3. rewrite H3. now apply H2. now apply H. Qed. End contents. math-classes-8.19.0/theory/forget_variety.v000066400000000000000000000012371460576051100207370ustar00rootroot00000000000000(* "Forgetting" a variety's laws (but keeping the algebraic operations) is a trivial functor. *) Require Import MathClasses.interfaces.canonical_names MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.interfaces.functors MathClasses.theory.categories MathClasses.categories.varieties MathClasses.categories.algebras. Section contents. Variable et: EquationalTheory. Definition forget (v: varieties.Object et) := algebras.object et v. Global Instance: Fmap forget := λ _ _, id. Global Instance: Functor forget _. Proof. constructor; intros; try apply _; repeat intro; try reflexivity. Qed. End contents. math-classes-8.19.0/theory/functors.v000066400000000000000000000034111460576051100175450ustar00rootroot00000000000000Require Import MathClasses.interfaces.canonical_names MathClasses.interfaces.abstract_algebra MathClasses.interfaces.functors. Require MathClasses.categories.setoids. Section setoid_functor_as_posh_functor. Context `{Pfunctor : @SFunctor F map_eq map}. Global Instance sfmap_mor `{Setoid_Morphism A B f} : Setoid_Morphism (sfmap f). Proof. pose proof (setoidmor_a f). pose proof (setoidmor_b f). split; try apply _. Qed. Lemma sfmap_id_applied `{Setoid A} (x : F A) : sfmap id x = x. Proof. change (sfmap id x = id x). apply setoids.ext_equiv_applied, sfmap_id. Qed. Lemma sfmap_comp_applied `{Equiv A} `{Equiv B} `{Equiv C} `{!Setoid_Morphism (f : B → C)} `{!Setoid_Morphism (g : A → B)} (x : F A) : sfmap (f ∘ g) x = sfmap f (sfmap g x). Proof. change (sfmap (f ∘ g) x = (sfmap f ∘ sfmap g) x). pose proof (setoidmor_a g). apply setoids.ext_equiv_applied. now apply sfmap_comp. Qed. Program Definition map_setoid: setoids.Object → setoids.Object := λ o, @setoids.object (F (setoids.T o)) (map_eq (setoids.T o) (setoids.e o)) _. Program Instance: Fmap map_setoid := λ v w x, @sfmap F map _ _ (`x). Instance: Functor map_setoid _. Proof. split; try apply _. intros [???] [???]. split; try apply _. intros [x ?] [y ?] ??? E1. simpl in *. assert (x = y) as E2 by intuition. now rewrite E1, E2. intros [???] x y E. simpl in *. now rewrite sfmap_id_applied. intros [???] [???] [??] [???] [??] ?? E. simpl in *. now rewrite sfmap_comp_applied, E. Qed. End setoid_functor_as_posh_functor. (** Note that we cannot prove the reverse (i.e. that an endo-Functor on setoid.Object gives rise to an SFunctor, because an SFunctor requires a Type→Type map, which we cannot get from a setoid.Object→setoid.Object map. *) math-classes-8.19.0/theory/groups.v000066400000000000000000000157531460576051100172350ustar00rootroot00000000000000Require MathClasses.theory.setoids. Require Import Coq.Classes.Morphisms MathClasses.interfaces.abstract_algebra. Section semigroup_props. Context `{SemiGroup G}. Global Instance sg_op_mor_1: ∀ x, Setoid_Morphism (x &) | 0. Proof. split; try apply _. Qed. Global Instance sg_op_mor_2: ∀ x, Setoid_Morphism (& x) | 0. Proof. split; try apply _. solve_proper. Qed. End semigroup_props. (** * This tactic simplifies group expressions by first pushing group morphisms and * negations down to the leaves, then cancelling inverses and removing units. *) Ltac group_simplify := rewrite_strat (try bottomup (hints group_simplify)); (try bottomup (choice (hints group_cancellation) ( <- associativity ))). Ltac group := group_simplify; easy. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Hint Rewrite @associativity using apply _: group_simplify. #[global] Hint Rewrite @left_identity @right_identity @left_inverse @right_inverse using apply _: group_cancellation. Set Warnings "+unsupported-attributes". Section group_props. Context `{Group G}. Lemma negate_sg_op_distr_flip `{Group A}: ∀ a b, -(a & b) = -b & -a. Proof. intros. setoid_replace (-b & -a) with (-(a & b) & (a & b) & (-b & -a)) by group; rewrite <- associativity; setoid_replace (a & b & (-b & -a)) with mon_unit; group. Qed. Hint Rewrite @negate_sg_op_distr_flip using apply _ : group_simplify. Global Instance negate_involutive: Involutive (-). Proof. intros x; setoid_replace x with (- - x & (- x & x)) at 2 by group; setoid_replace (-x & x) with mon_unit by group; group. Qed. Hint Rewrite @negate_involutive using apply _ : group_cancellation. Global Instance: Injective (-). Proof. repeat (split; try apply _). intros x y E. now rewrite <-(involutive x), <-(involutive y), E. Qed. Lemma negate_mon_unit : -mon_unit = mon_unit. Proof. setoid_replace mon_unit with (-mon_unit & mon_unit); group. Qed. Hint Rewrite @negate_mon_unit using apply _ : group_cancellation. Global Instance: ∀ z : G, LeftCancellation (&) z. Proof. intros z x y E; setoid_replace x with (-z & (z & x)) by group; rewrite E; group. Qed. Global Instance: ∀ z : G, RightCancellation (&) z. Proof. intros z x y E; setoid_replace x with (x & z & -z) by group; rewrite E; group. Qed. Lemma negate_sg_op_distr `{!AbGroup G} x y: -(x & y) = -x & -y. Proof. group_simplify. apply commutativity. Qed. End group_props. Section groupmor_props. Context `{Group A} `{Group B} {f : A → B} `{!Monoid_Morphism f}. Lemma preserves_negate x : f (-x) = -f x. Proof. apply (left_cancellation (&) (f x)); rewrite <- preserves_sg_op; group_simplify; apply preserves_mon_unit. Qed. Hint Rewrite @preserves_sg_op @preserves_negate @preserves_mon_unit using apply _ : group_simplify. End groupmor_props. #[global] Instance semigroup_morphism_proper {A B eA eB opA opB} : Proper ((=) ==> iff) (@SemiGroup_Morphism A B eA eB opA opB) | 1. Proof. assert (∀ (f g : A → B), g = f → SemiGroup_Morphism f → SemiGroup_Morphism g) as P. intros f g E [? ? ? ?]. split; try apply _. eapply setoids.morphism_proper; eauto. intros x y. now rewrite (E (x & y)), (E x), (E y). intros f g ?; split; intros Mor. apply P with f. destruct Mor. now symmetry. apply _. now apply P with g. Qed. #[global] Instance monoid_morphism_proper {A B eA eB opA uA opB uB} : Proper ((=) ==> iff) (@Monoid_Morphism A B eA eB opA uA opB uB) | 1. Proof. assert (∀ (f g : A → B), g = f → Monoid_Morphism f → Monoid_Morphism g) as P. intros f g E [? ? ? ?]. split; try apply _. eapply semigroup_morphism_proper; eauto. now rewrite (E mon_unit mon_unit). intros f g ?; split; intros Mor. apply P with f. destruct Mor. now symmetry. apply _. now apply P with g. Qed. Section from_another_sg. Context `{SemiGroup A} `{Setoid B} `{Bop : SgOp B} (f : B → A) `{!Injective f} (op_correct : ∀ x y, f (x & y) = f x & f y). Instance: Setoid_Morphism f := injective_mor f. Instance: Proper ((=) ==> (=) ==> (=)) Bop. Proof. intros ? ? E1 ? ? E2. apply (injective f). rewrite 2!op_correct. apply sg_op_proper; now apply sm_proper. Qed. Lemma projected_sg: SemiGroup B. Proof. split; try apply _. repeat intro; apply (injective f). now rewrite !op_correct, associativity. Qed. End from_another_sg. Section from_another_com_sg. Context `{CommutativeSemiGroup A} `{Setoid B} `{Bop : SgOp B} (f : B → A) `{!Injective f} (op_correct : ∀ x y, f (x & y) = f x & f y). Lemma projected_com_sg: CommutativeSemiGroup B. Proof. split. now apply (projected_sg f). repeat intro; apply (injective f). now rewrite !op_correct, commutativity. Qed. End from_another_com_sg. Section from_another_monoid. Context `{Monoid A} `{Setoid B} `{Bop : SgOp B} `{Bunit : MonUnit B} (f : B → A) `{!Injective f} (op_correct : ∀ x y, f (x & y) = f x & f y) (unit_correct : f mon_unit = mon_unit). Lemma projected_monoid: Monoid B. Proof. split. now apply (projected_sg f). repeat intro; apply (injective f). now rewrite op_correct, unit_correct, left_identity. repeat intro; apply (injective f). now rewrite op_correct, unit_correct, right_identity. Qed. End from_another_monoid. Section from_another_com_monoid. Context `{CommutativeMonoid A} `{Setoid B} `{Bop : SgOp B} `{Bunit : MonUnit B} (f : B → A) `{!Injective f} (op_correct : ∀ x y, f (x & y) = f x & f y) (unit_correct : f mon_unit = mon_unit). Lemma projected_com_monoid: CommutativeMonoid B. Proof. split. now apply (projected_monoid f). repeat intro; apply (injective f). now rewrite op_correct, commutativity. Qed. End from_another_com_monoid. Section from_another_group. Context `{Group A} `{Setoid B} `{Bop : SgOp B} `{Bunit : MonUnit B} `{Bnegate : Negate B} (f : B → A) `{!Injective f} (op_correct : ∀ x y, f (x & y) = f x & f y) (unit_correct : f mon_unit = mon_unit) (negate_correct : ∀ x, f (-x) = -f x). Instance: Setoid_Morphism f := injective_mor f. Instance: Setoid_Morphism Bnegate. Proof. split; try apply _. intros ? ? E1. apply (injective f). rewrite 2!negate_correct. now do 2 apply sm_proper. Qed. Lemma projected_group: Group B. Proof. split; try apply _. now apply (projected_monoid f). repeat intro; apply (injective f). now rewrite op_correct, negate_correct, unit_correct, left_inverse. repeat intro; apply (injective f). now rewrite op_correct, negate_correct, unit_correct, right_inverse. Qed. End from_another_group. Section from_another_ab_group. Context `{AbGroup A} `{Setoid B} `{Bop : SgOp B} `{Bunit : MonUnit B} `{Bnegate : Negate B} (f : B → A) `{!Injective f} (op_correct : ∀ x y, f (x & y) = f x & f y) (unit_correct : f mon_unit = mon_unit) (negate_correct : ∀ x, f (-x) = -f x). Lemma projected_ab_group: AbGroup B. Proof. split; try apply _. now apply (projected_group f). repeat intro; apply (injective f). now rewrite op_correct, commutativity. Qed. End from_another_ab_group. math-classes-8.19.0/theory/hom_functor.v000066400000000000000000000014411460576051100202260ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.theory.setoids MathClasses.interfaces.functors MathClasses.theory.categories. Require MathClasses.categories.setoids. Section contents. Context `{Category C} (x: C). Definition homFrom (y: C): setoids.Object := setoids.object (x ⟶ y). Global Program Instance: Fmap homFrom := λ v w X, (X ◎): (x ⟶ v) → (x ⟶ w). Next Obligation. constructor; apply _. Qed. Global Instance: Functor homFrom _. Proof. constructor; try apply _. constructor; try apply _. repeat intro. simpl in *. apply comp_proper; auto. repeat intro. simpl. rewrite H1. apply left_identity. repeat intro. simpl. unfold compose. rewrite H1. symmetry. apply comp_assoc. Qed. End contents. math-classes-8.19.0/theory/int_abs.v000066400000000000000000000102331460576051100173210ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.naturals MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.orders.nat_int MathClasses.theory.integers MathClasses.theory.rings MathClasses.orders.rings. Section contents. Context `{Integers Z} `{Apart Z} `{!TrivialApart Z} `{!FullPseudoSemiRingOrder Zle Zlt} `{Naturals N}. Add Ring Z : (rings.stdlib_ring_theory Z). Lemma int_abs_unique_respectful {a b : IntAbs Z N} : ((=) ==> (=))%signature (int_abs Z N (ia:=a)) (int_abs Z N (ia:= b)). Proof. intros x y E. unfold int_abs, int_abs_sig. apply (injective (naturals_to_semiring N Z)). destruct a as [[z1 A]|[z1 A]], b as [[z2 B]|[z2 B]]. now rewrite A, B. symmetry. apply naturals.negate_to_ring. now rewrite A, B, involutive. apply naturals.negate_to_ring. now rewrite A, B, involutive. apply (injective (-)). now rewrite A, B, !involutive. Qed. Lemma int_abs_unique (a b : IntAbs Z N) (z : Z) : int_abs Z N (ia:=a) z = int_abs Z N (ia:=a) z. Proof. now apply int_abs_unique_respectful. Qed. Context `{!IntAbs Z N}. Global Instance int_abs_proper: Setoid_Morphism (int_abs Z N) | 0. Proof. split; try apply _. now apply int_abs_unique_respectful. Qed. Context `{!SemiRing_Morphism (f : N → Z)}. Lemma int_abs_spec x : { 0 ≤ x ∧ f (int_abs Z N x) = x } + { x ≤ 0 ∧ f (int_abs Z N x) = -x }. Proof. unfold int_abs. destruct int_abs_sig as [[n E]|[n E]]. left. rewrite <-E. split. now apply to_semiring_nonneg. apply (naturals.to_semiring_unique_alt _ _). right. split. apply flip_nonpos_negate. rewrite <-E. now apply to_semiring_nonneg. rewrite <-E. now apply (naturals.to_semiring_unique_alt _ _). Qed. Lemma int_abs_sig_alt x : { n : N | f n = x } + { n : N | f n = -x }. Proof. destruct (int_abs_spec x) as [[??]|[??]]; eauto. Qed. Lemma int_abs_nat n : int_abs Z N (f n) = n. Proof. apply (injective f). destruct (int_abs_spec (f n)) as [[? E]|[? E]]; intuition. apply naturals.negate_to_ring. now rewrite E, involutive. Qed. Lemma int_abs_negate_nat n : int_abs Z N (-f n) = n. Proof. apply (injective f). destruct (int_abs_spec (-f n)) as [[? E]|[? E]]. symmetry. now apply naturals.negate_to_ring. now rewrite involutive in E. Qed. Lemma int_abs_negate x : int_abs Z N (-x) = int_abs Z N x. Proof. destruct (int_abs_spec x) as [[_ E]|[_ E]]. rewrite <-E at 1. now apply int_abs_negate_nat. rewrite <-E at 1. now apply int_abs_nat. Qed. Lemma int_abs_0_alt x : int_abs Z N x = 0 ↔ x = 0. Proof. split; intros E1. destruct (int_abs_spec x) as [[_ E2]|[_ E2]]. now rewrite <-E2, E1, preserves_0. apply flip_negate_0. now rewrite <-E2, E1, preserves_0. rewrite E1, <-(preserves_0 (f:=f)). now apply int_abs_nat. Qed. Lemma int_abs_ne_0 x : int_abs Z N x ≠ 0 ↔ x ≠ 0. Proof. destruct (int_abs_0_alt x). intuition. Defined. Lemma int_abs_0 : int_abs Z N 0 = 0. Proof. now apply int_abs_0_alt. Qed. Lemma int_abs_nonneg x : 0 ≤ x → f (int_abs Z N x) = x. Proof. intros E1. destruct (int_abs_spec x); intuition. setoid_replace x with (0 : Z). now rewrite int_abs_0, preserves_0. now apply (antisymmetry (≤)). Qed. Lemma int_abs_nonpos x : x ≤ 0 → f (int_abs Z N x) = -x. Proof. intros E. rewrite <-int_abs_negate, int_abs_nonneg; intuition. now apply flip_nonpos_negate. Qed. Lemma int_abs_1 : int_abs Z N 1 = 1. Proof. apply (injective f). rewrite preserves_1. apply int_abs_nonneg; solve_propholds. Qed. Lemma int_abs_nonneg_plus x y : 0 ≤ x → 0 ≤ y → int_abs Z N (x + y) = int_abs Z N x + int_abs Z N y. Proof. intros. apply (injective f). rewrite preserves_plus, !int_abs_nonneg; intuition. now apply nonneg_plus_compat. Qed. Lemma int_abs_mult x y : int_abs Z N (x * y) = int_abs Z N x * int_abs Z N y. Proof. apply (injective f). rewrite preserves_mult. destruct (int_abs_spec x) as [[? Ex]|[? Ex]], (int_abs_spec y) as [[? Ey]|[? Ey]]; rewrite Ex, Ey. rewrite int_abs_nonneg. easy. now apply nonneg_mult_compat. rewrite int_abs_nonpos. ring. now apply nonneg_nonpos_mult. rewrite int_abs_nonpos. ring. now apply nonpos_nonneg_mult. rewrite int_abs_nonneg. ring. now apply nonpos_mult. Qed. End contents. math-classes-8.19.0/theory/int_pow.v000066400000000000000000000344241460576051100173710ustar00rootroot00000000000000Require MathClasses.theory.naturals MathClasses.orders.semirings MathClasses.orders.integers MathClasses.orders.dec_fields. Require Import Coq.setoid_ring.Ring Coq.setoid_ring.Field MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals MathClasses.interfaces.integers MathClasses.interfaces.additional_operations MathClasses.interfaces.orders MathClasses.theory.nat_pow MathClasses.theory.int_abs MathClasses.theory.dec_fields. (* * Properties of Int Pow *) Section int_pow_properties. Context `{DecField A} `{∀ x y, Decision (x = y)} `{Integers B} `{!IntPowSpec A B ipw}. Add Field A : (dec_fields.stdlib_field_theory A). Add Ring B : (rings.stdlib_ring_theory B). Global Instance: Proper ((=) ==> (=) ==> (=)) ((^) : A → B → A) | 0. Proof int_pow_proper. Global Instance int_pow_mor_1: ∀ x : A, Setoid_Morphism (x^) | 0. Proof. split; try apply _. Qed. Global Instance int_pow_mor_2: ∀ n : B, Setoid_Morphism (^n) | 0. Proof. split; try apply _. solve_proper. Qed. Lemma int_pow_S_nonneg `{Apart B} `{!TrivialApart B} `{!FullPseudoSemiRingOrder (A:=B) Ble Blt} (x : A) (n : B) : 0 ≤ n → x ^ (1+n) = x * x ^ n. Proof. intros En. destruct (decide (x = 0)) as [Ex | Ex]. rewrite Ex. rewrite int_pow_base_0. ring. intros E. destruct semirings.not_le_1_0. rewrite <-E. now apply semirings.nonneg_plus_le_compat_r. now rewrite int_pow_S. Qed. Lemma int_pow_negate (x : A) (n : B) : x ^ (-n) = /(x ^ n). Proof. destruct (decide (x = 0)) as [Ex | Ex]. rewrite Ex. destruct (decide (n = 0)) as [En | En]. now rewrite En, rings.negate_0, int_pow_0, dec_recip_1. rewrite 2!int_pow_base_0; trivial. now rewrite dec_recip_0. now apply rings.flip_negate_ne_0. revert n. apply biinduction. solve_proper. now rewrite rings.negate_0, int_pow_0, dec_recip_1. intros n. setoid_replace (-n) with (1 - (1 + n)) by ring. rewrite 2!int_pow_S, dec_recip_distr; trivial. split; intros E. rewrite <-E. now field. rewrite E, associativity, dec_recip_inverse; trivial. ring. Qed. Lemma int_pow_negate_alt (x : A) (n : B) : x ^ n = /(x ^ (-n)). Proof. rewrite <-int_pow_negate. now rewrite rings.negate_involutive. Qed. Lemma int_pow_mult (x y : A) (n : B) : (x * y) ^ n = x ^ n * y ^ n. Proof. destruct (decide (x * y = 0)) as [Exy | Exy]. rewrite Exy. destruct (decide (n = 0)) as [En | En]. rewrite En, 3!int_pow_0. ring. destruct (zero_product x y Exy) as [E|E]; rewrite E, int_pow_base_0; trivial; ring. revert n. apply biinduction. solve_proper. rewrite 3!int_pow_0. ring. intros n. rewrite 3!int_pow_S; trivial. split; intros E. rewrite E. ring. apply (rings.left_cancellation_ne_0 (.*.) (x * y)); trivial. rewrite E. ring. intros E. apply Exy. rewrite E. ring. intros E. apply Exy. rewrite E. ring. Qed. Lemma int_pow_recip (x : A) (n : B) : (/x) ^ n = /(x ^ n). Proof. destruct (decide (x = 0)) as [Ex | Ex]. rewrite Ex, dec_recip_0. destruct (decide (n = 0)) as [En | En]. now rewrite En, int_pow_0, dec_recip_1. now rewrite int_pow_base_0, dec_recip_0. revert n. apply biinduction. solve_proper. now rewrite 2!int_pow_0, dec_recip_1. intros n. assert (/x ≠ 0) by now apply dec_recip_ne_0. rewrite 2!int_pow_S, dec_recip_distr; trivial. split; intros E. now rewrite E. now apply (rings.left_cancellation_ne_0 (.*.) (/x)). Qed. Lemma int_pow_nat_pow `{Naturals N} `{!NatPowSpec A N pw} {f : N → B} `{!SemiRing_Morphism f} (x : A) (n : N) : x ^ (f n) = x ^ n. Proof. revert n. apply naturals.induction. solve_proper. now rewrite rings.preserves_0, int_pow_0, nat_pow_0. intros n E. rewrite rings.preserves_plus, rings.preserves_1. rewrite int_pow_S_nonneg, nat_pow_S. now rewrite E. apply semirings.preserves_nonneg. now apply naturals.nat_nonneg. Qed. Global Instance int_pow_1: RightIdentity (^) (1:B). Proof. intro. assert ((1:B) = 1 + 0) as E by ring. rewrite E. rewrite int_pow_S_nonneg, int_pow_0; [ring | reflexivity]. Qed. Lemma int_pow_2 x : x ^ (2:B) = x * x. Proof. now rewrite int_pow_S_nonneg, int_pow_1 by solve_propholds. Qed. Lemma int_pow_3 x : x ^ (3:B) = x * (x * x). Proof. now rewrite int_pow_S_nonneg, int_pow_2 by solve_propholds. Qed. Lemma int_pow_4 x : x ^ (4:B) = x * (x * (x * x)). Proof. now rewrite int_pow_S_nonneg, int_pow_3 by solve_propholds. Qed. Global Instance int_pow_base_1: LeftAbsorb (^) (1:A). Proof. red. apply biinduction. solve_proper. now apply int_pow_0. intros n. rewrite int_pow_S, left_identity. easy. now apply (rings.is_ne_0 1). Qed. Lemma int_pow_exp_plus (n m : B) (x : A) : x ≠ 0 → x ^ (n + m) = x ^ n * x ^ m. Proof. intros nonneg. revert n. apply biinduction. solve_proper. rewrite int_pow_0, left_identity. ring. intros n. rewrite <-associativity, 2!int_pow_S; trivial. split; intros E. rewrite E. ring. apply (rings.left_cancellation_ne_0 (.*.) x); trivial. rewrite E. ring. Qed. Instance int_pow_ne_0 (x : A) (n : B) : PropHolds (x ≠ 0) → PropHolds (x ^ n ≠ 0). Proof. intros nonneg. unfold PropHolds. revert n. apply biinduction. solve_proper. rewrite int_pow_0. apply (rings.is_ne_0 1). intros n. rewrite int_pow_S; trivial. split; intros E1 E2; destruct E1. apply (left_cancellation (.*.) x). now rewrite right_absorb. rewrite E2. ring. Qed. Lemma int_pow_exp_mult (x : A) (n m : B) : x ^ (n * m) = (x ^ n) ^ m. Proof. destruct (decide (x = 0)) as [Ex|Ex]. rewrite Ex. destruct (decide (n = 0)) as [En|En]. rewrite En, left_absorb, int_pow_0. now rewrite left_absorb. destruct (decide (m = 0)) as [Em|Em]. now rewrite Em, right_absorb, 2!int_pow_0. rewrite 3!int_pow_base_0; try easy. intros E. now destruct (zero_product n m E). revert m. apply biinduction. solve_proper. rewrite right_absorb. now rewrite 2!int_pow_0. intros m. split; intros E. rewrite int_pow_S, <-E. rewrite distribute_l, right_identity. now rewrite int_pow_exp_plus. now apply int_pow_ne_0. rewrite int_pow_S in E. rewrite distribute_l, right_identity, int_pow_exp_plus in E. apply (rings.left_cancellation_ne_0 (.*.) (x ^ n)). now apply int_pow_ne_0. now rewrite E. easy. now apply int_pow_ne_0. Qed. Context `{Apart A} `{!TrivialApart A} `{!FullPseudoSemiRingOrder (A:=A) Ale Alt}. Context `{Apart B} `{!TrivialApart B} `{!FullPseudoSemiRingOrder (A:=B) Ble Blt}. Instance int_pow_pos (x : A) (n : B) : PropHolds (0 < x) → PropHolds (0 < x ^ n). Proof. intros nonneg. unfold PropHolds. revert n. apply biinduction. solve_proper. intros. rewrite int_pow_0. now apply semirings.lt_0_1. intros n; split; intros E. rewrite int_pow_S. now apply pos_mult_compat. apply not_symmetry. now apply orders.lt_ne. apply (strictly_order_reflecting (x *.)). rewrite <-int_pow_S. now rewrite right_absorb. apply not_symmetry. now apply orders.lt_ne. Qed. Instance int_pow_nonneg (x : A) (n : B) : PropHolds (0 ≤ x) → PropHolds (0 ≤ x ^ n). Proof. intros E1. red in E1. destruct (orders.le_equiv_lt _ _ E1) as [E2|E2]. rewrite <-E2. destruct (decide (n = 0)) as [En|En]. rewrite En. rewrite int_pow_0. apply semirings.le_0_1. unfold PropHolds. now rewrite int_pow_base_0. now apply orders.lt_le, int_pow_pos. Qed. Lemma int_pow_ge_1 (x : A) (n : B) : 1 ≤ x → 0 ≤ n → 1 ≤ x ^ n. Proof. intros E1 E2. revert n E2. apply integers.induction_nonneg; trivial. solve_proper. now rewrite int_pow_0. intros. rewrite int_pow_S. rewrite <-rings.mult_1_r. apply semirings.mult_le_compat; try apply semirings.le_0_1; auto. apply orders.lt_ne_flip. apply orders.lt_le_trans with 1; trivial. now apply semirings.lt_0_1. Qed. Lemma int_pow_gt_1 (x : A) (n : B) : 1 < x → 0 < n → 1 < x ^ n. Proof. intros Ex En. apply nat_int.lt_iff_S_le in En. destruct (semirings.decompose_le En) as [z [Ez1 Ez2]]. ring_simplify in Ez2. rewrite Ez2. clear En Ez2 n. revert z Ez1. apply integers.induction_nonneg; try assumption. solve_proper. now rewrite left_identity, right_identity. intros n En E2. rewrite <-associativity, int_pow_S. apply semirings.gt_1_mult_lt_compat_l; auto. transitivity (1:A); [apply semirings.lt_0_1 | assumption]. apply orders.lt_ne_flip. apply orders.le_lt_trans with 1; trivial. now apply semirings.le_0_1. Qed. (* Making these instances Global is not useful, we don't have PropHolds (1 ≤ x) instances and it will only slow down instance resolution (it increases the compilation time of dyadics from 1:35 to 2:28). *) Instance int_pow_exp_le: ∀ x : A, PropHolds (1 ≤ x) → OrderPreserving (x^). Proof. repeat (split; try apply _). assert (0 < x) by (apply orders.lt_le_trans with 1; [solve_propholds | easy]). intros n m E. destruct (semirings.decompose_le E) as [z [Ea Eb]]. rewrite Eb. rewrite int_pow_exp_plus by now apply orders.lt_ne_flip. rewrite <-rings.mult_1_r at 1. apply (order_preserving (x ^ n *.)). now apply int_pow_ge_1. Qed. Instance int_pow_exp_lt: ∀ x : A, PropHolds (1 < x) → StrictlyOrderPreserving (x^). Proof. repeat (split; try apply _). assert (0 < x) by (apply orders.le_lt_trans with 1; [solve_propholds | easy]). intros n m E. apply nat_int.lt_iff_plus_1_le in E. destruct (semirings.decompose_le E) as [z [Ea Eb]]. rewrite Eb. rewrite <-associativity, int_pow_exp_plus by now apply orders.lt_ne_flip. rewrite <-(rings.mult_1_r (x ^ n)) at 1. apply (strictly_order_preserving (x^n *.)). apply int_pow_gt_1; trivial. now apply nat_int.le_iff_lt_S. Qed. Instance int_pow_exp_le_back: ∀ x : A, PropHolds (1 < x) → OrderReflecting (x^). Proof. split; try apply _. intros n m E1. destruct (total (≤) n m) as [E2|E2]; trivial. destruct (orders.le_equiv_lt _ _ E2) as [E3|E3]. now rewrite E3. contradict E1. apply orders.lt_not_le_flip. now apply (strictly_order_preserving (x^)). Qed. Instance int_pow_exp_lt_back: ∀ x : A, PropHolds (1 < x) → StrictlyOrderReflecting (x^). Proof. intros ? E1. apply _. Qed. Instance int_pow_inj: ∀ x : A, PropHolds (1 < x) → Injective (x^). Proof. repeat (split; try apply _). intros n m E. apply (antisymmetry (≤)); apply (order_reflecting (x^)); trivial; rewrite E; reflexivity. Qed. End int_pow_properties. (* Due to bug #2528 *) #[global] Hint Extern 18 (PropHolds (_ ^ _ ≠ 0)) => eapply @int_pow_ne_0 : typeclass_instances. #[global] Hint Extern 18 (PropHolds (0 ≤ _ ^ _)) => eapply @int_pow_nonneg : typeclass_instances. #[global] Hint Extern 18 (PropHolds (0 < _ ^ _)) => eapply @int_pow_pos : typeclass_instances. Section preservation. Context `{Integers B} `{DecField A1} `{∀ x y : A1, Decision (x = y)} `{!IntPowSpec A1 B ip1} `{DecField A2} `{∀ x y : A2, Decision (x = y)} `{!IntPowSpec A2 B ip2} {f : A1 → A2} `{!SemiRing_Morphism f}. Add Ring B2 : (rings.stdlib_ring_theory B). Lemma preserves_int_pow x (n : B) : f (x ^ n) = (f x) ^ n. Proof. destruct (decide (x = 0)) as [Ex | Ex]. rewrite Ex, rings.preserves_0. destruct (decide (n = 0)) as [En|En]. rewrite En, 2!int_pow_0. now apply rings.preserves_1. rewrite 2!int_pow_base_0; trivial. now apply rings.preserves_0. revert n. apply biinduction. solve_proper. rewrite int_pow_0, int_pow_0. now apply rings.preserves_1. intros n. assert (f x ≠ 0) by now apply rings.injective_ne_0. rewrite 2!int_pow_S, rings.preserves_mult; trivial. split; intros E. now rewrite E. now apply (left_cancellation (.*.) (f x)). Qed. End preservation. Section exp_preservation. Context `{Field A} `{∀ x y : A, Decision (x = y)} `{Integers B1} `{Integers B2} `{!IntPowSpec A B1 pw1} `{!IntPowSpec A B2 pw2} {f : B1 → B2} `{!SemiRing_Morphism f}. Lemma preserves_int_pow_exp x (n : B1) : x ^ (f n) = x ^ n. Proof. destruct (decide (x = 0)) as [Ex | Ex]. rewrite Ex. destruct (decide (n = 0)) as [En|En]. now rewrite En, rings.preserves_0, 2!int_pow_0. rewrite 2!int_pow_base_0; try easy. now apply rings.injective_ne_0. revert n. apply biinduction. solve_proper. rewrite rings.preserves_0. now rewrite 2!int_pow_0. intros n. rewrite rings.preserves_plus, rings.preserves_1. rewrite 2!int_pow_S by trivial. split; intros E. now rewrite E. now apply (rings.left_cancellation_ne_0 (.*.) x). Qed. End exp_preservation. (* Very slow default implementation by translation into Peano *) Section int_pow_default. Context `{DecField A} `{∀ x y, Decision (x = y)} `{Integers B} `{Apart B} `{!TrivialApart B} `{!FullPseudoSemiRingOrder (A:=B) Ble Blt}. Add Ring B3 : (rings.stdlib_ring_theory B). Global Instance int_pow_default: Pow A B | 10 := λ x n, match (decide_rel (≤) 0 n) with | left _ => x ^ int_abs B nat n | right _ => /x ^ int_abs B nat n end. Global Instance: IntPowSpec A B int_pow_default. Proof. split; unfold pow, int_pow_default. intros ? ? E1 ? ? E2. now (case (decide_rel); case (decide_rel); rewrite E1, E2). intros x. case (decide_rel); intros E. now rewrite int_abs_0. now destruct E. intros n ?. case (decide_rel); intros E. now apply nat_pow_base_0, int_abs_ne_0. rewrite nat_pow_base_0. apply dec_recip_0. now apply int_abs_ne_0. intros x n E. case (decide_rel); case (decide_rel); intros E1 E2. now rewrite int_abs_nonneg_plus, int_abs_1 by (auto;solve_propholds). setoid_replace n with (-1 : B). rewrite rings.plus_negate_r, int_abs_0, nat_pow_0. rewrite int_abs_negate, int_abs_1, right_identity. symmetry. now apply dec_recip_inverse. apply (antisymmetry (≤)). apply orders.not_le_lt_flip in E1. apply nat_int.lt_iff_plus_1_le in E1. apply (order_reflecting (+1)). now ring_simplify. apply (order_reflecting (1+)). now rewrite rings.plus_negate_r. destruct E2. apply semirings.nonneg_plus_compat; [solve_propholds | assumption]. rewrite <-int_abs_negate, <-(int_abs_negate n). setoid_replace (-n) with (1 - (1 + n)) by ring. rewrite (int_abs_nonneg_plus 1 (-(1 + n))), int_abs_1. rewrite nat_pow_S. rewrite dec_recip_distr, associativity. now rewrite dec_recip_inverse, left_identity. now apply (rings.is_nonneg 1). now apply rings.flip_nonpos_negate, orders.le_flip. Qed. End int_pow_default. math-classes-8.19.0/theory/int_to_nat.v000066400000000000000000000156671460576051100200600ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.naturals MathClasses.interfaces.abstract_algebra MathClasses.interfaces.orders MathClasses.orders.nat_int MathClasses.theory.integers MathClasses.theory.rings MathClasses.orders.rings. Section contents. Context `{Integers Z} `{Apart Z} `{!TrivialApart Z} `{!FullPseudoSemiRingOrder Zle Zlt} `{Naturals N}. Add Ring Z : (rings.stdlib_ring_theory Z). Lemma int_to_nat_unique_respectful {a b : IntAbs Z N} : ((=) ==> (=))%signature (int_to_nat Z N (ia:=a)) (int_to_nat Z N (ia:= b)). Proof. intros x y E. unfold int_to_nat, int_abs_sig. apply (injective (naturals_to_semiring N Z)). destruct a as [[z1 A]|[z1 A]], b as [[z2 B]|[z2 B]]. now rewrite A, B. destruct (naturals.to_ring_zero_sum z2 z1) as [? E2]. now rewrite B, A, involutive. now rewrite E2. destruct (naturals.to_ring_zero_sum z1 z2) as [? E2]. now rewrite B, A, involutive. now rewrite E2. reflexivity. Qed. Lemma int_to_nat_unique (a b : IntAbs Z N) (z : Z) : int_to_nat Z N (ia:=a) z = int_to_nat Z N (ia:=a) z. Proof. now apply int_to_nat_unique_respectful. Qed. Context `{!IntAbs Z N}. Global Instance int_to_nat_proper: Setoid_Morphism (int_to_nat Z N) | 0. Proof. split; try apply _. now apply int_to_nat_unique_respectful. Qed. Context `{!SemiRing_Morphism (f : N → Z)}. Lemma int_to_nat_spec x : { 0 ≤ x ∧ f (int_to_nat Z N x) = x } + { x ≤ 0 ∧ int_to_nat Z N x = 0 }. Proof. unfold int_to_nat. destruct int_abs_sig as [[n E]|[n E]]. left. rewrite <-E. split. now apply to_semiring_nonneg. apply (naturals.to_semiring_unique_alt _ _). right. intuition. apply flip_nonpos_negate. rewrite <-E. now apply to_semiring_nonneg. Qed. Lemma int_to_nat_nat n : int_to_nat Z N (f n) = n. Proof. apply (injective f). destruct (int_to_nat_spec (f n)) as [[??]|[? E]]; intuition. rewrite E, preserves_0. apply (antisymmetry (≤)); intuition. now apply to_semiring_nonneg. Qed. Lemma int_to_nat_cancel_l x n : x = f n → int_to_nat Z N x = n. Proof. intros E. now rewrite E, int_to_nat_nat. Qed. Lemma int_to_nat_cancel_r x n : f n = x → n = int_to_nat Z N x. Proof. intros E. now rewrite <-E, int_to_nat_nat. Qed. Lemma int_to_nat_negate_nat n : int_to_nat Z N (-f n) = 0. Proof. apply (injective f). destruct (int_to_nat_spec (-f n)) as [[? E]|[? E]]. rewrite E, preserves_0. apply (antisymmetry (≤)); intuition. apply rings.flip_nonneg_negate. now apply to_semiring_nonneg. now rewrite E. Qed. Lemma int_to_nat_0 : int_to_nat Z N 0 = 0. Proof. apply int_to_nat_cancel_l. now rewrite preserves_0. Qed. Lemma int_to_nat_of_nonneg x : 0 ≤ x → f (int_to_nat Z N x) = x. Proof. intros E1. destruct (int_to_nat_spec x) as [[? E2]|[? E2]]; intuition. rewrite E2, preserves_0. now apply (antisymmetry (≤)). Qed. Lemma int_to_nat_of_nonpos x : x ≤ 0 → int_to_nat Z N x = 0. Proof. intros E. destruct (int_to_nat_spec x) as [[? E2]|]; intuition. apply int_to_nat_cancel_l. rewrite preserves_0. now apply (antisymmetry (≤)). Qed. Lemma int_to_nat_nonneg_plus x y : 0 ≤ x → 0 ≤ y → int_to_nat Z N (x + y) = int_to_nat Z N x + int_to_nat Z N y. Proof. intros. apply (injective f). rewrite preserves_plus, !int_to_nat_of_nonneg; intuition. now apply nonneg_plus_compat. Qed. Lemma int_to_nat_mult_nonneg_l x y : 0 ≤ x → int_to_nat Z N (x * y) = int_to_nat Z N x * int_to_nat Z N y. Proof. intros E. apply (injective f). rewrite preserves_mult. rewrite (int_to_nat_of_nonneg x) by easy. destruct (int_to_nat_spec y) as [[? Ey]|[? Ey]]; rewrite Ey, ?preserves_0. rewrite int_to_nat_of_nonneg. easy. now apply nonneg_mult_compat. rewrite int_to_nat_of_nonpos, preserves_0. ring. now apply nonneg_nonpos_mult. Qed. Lemma int_to_nat_mult_nonneg_r x y : 0 ≤ y → int_to_nat Z N (x * y) = int_to_nat Z N x * int_to_nat Z N y. Proof. rewrite (commutativity x), (commutativity (int_to_nat Z N x)). now apply int_to_nat_mult_nonneg_l. Qed. Lemma int_to_nat_1 : int_to_nat Z N 1 = 1. Proof. apply int_to_nat_cancel_l. now rewrite preserves_1. Qed. Context `{Apart N} `{!TrivialApart N} `{!FullPseudoSemiRingOrder (A:=N) Nle Nlt}. Global Instance int_to_nat_nonneg x : PropHolds (0 ≤ int_to_nat Z N x). Proof. destruct (int_to_nat_spec x) as [[? E]|[? E]]. apply (order_reflecting f). now rewrite preserves_0, E. rewrite E. solve_propholds. Qed. Global Instance int_to_nat_pos x : PropHolds (0 < x) → PropHolds (0 < int_to_nat Z N x). Proof. rewrite !lt_iff_le_ne. intros [??]. split. solve_propholds. apply (setoids.morphism_ne f). now rewrite preserves_0, int_to_nat_of_nonneg. Qed. Lemma int_to_nat_le_l x y : 0 ≤ x → x ≤ y → f (int_to_nat Z N x) ≤ y. Proof. intros. now rewrite int_to_nat_of_nonneg. Qed. Lemma int_to_nat_le_cancel_l x n : 0 ≤ x → x ≤ f n → int_to_nat Z N x ≤ n. Proof. intros. now apply (order_reflecting f), int_to_nat_le_l. Qed. Lemma int_to_nat_le_r x y : x ≤ y → x ≤ f (int_to_nat Z N y). Proof. intros E1. destruct (int_to_nat_spec y) as [[? E2]|[? E2]]. now rewrite E2. rewrite E2, preserves_0. now transitivity y. Qed. Lemma int_to_nat_le_cancel_r x n : f n ≤ x → n ≤ int_to_nat Z N x. Proof. intros. now apply (order_reflecting f), int_to_nat_le_r. Qed. Global Instance: OrderPreserving (int_to_nat Z N). Proof. repeat (split; try apply _). intros x y E. destruct (total (≤) 0 x). now apply int_to_nat_le_cancel_r, int_to_nat_le_l. rewrite int_to_nat_of_nonpos. solve_propholds. easy. Qed. Lemma int_to_nat_le_back x y : 0 ≤ y → int_to_nat Z N x ≤ int_to_nat Z N y → x ≤ y. Proof. intros. rewrite <-(int_to_nat_of_nonneg y) by easy. transitivity (f (int_to_nat Z N x)). now apply int_to_nat_le_r. now apply (order_preserving f). Qed. Lemma int_to_nat_lt_l x y : 0 ≤ x → x < y → f (int_to_nat Z N x) < y. Proof. intros. now rewrite int_to_nat_of_nonneg. Qed. Lemma int_to_nat_lt_cancel_l x n : 0 ≤ x → x < f n → int_to_nat Z N x < n. Proof. intros. now apply (strictly_order_reflecting f), int_to_nat_lt_l. Qed. Lemma int_to_nat_lt_r x y : x < y → x < f (int_to_nat Z N y). Proof. intros E1. destruct (int_to_nat_spec y) as [[? E2]|[? E2]]. now rewrite E2. rewrite E2, preserves_0. now apply lt_le_trans with y. Qed. Lemma int_to_nat_lt_cancel_r x n : f n < x → n < int_to_nat Z N x. Proof. intros. now apply (strictly_order_reflecting f), int_to_nat_lt_r. Qed. Lemma int_to_nat_lt x y : 0 < y → x < y → int_to_nat Z N x < int_to_nat Z N y. Proof. intros Ey Exy. destruct (total (≤) 0 x). now apply int_to_nat_lt_cancel_r, int_to_nat_lt_l. rewrite int_to_nat_of_nonpos by easy. now apply int_to_nat_pos. Qed. Lemma int_to_nat_lt_back x y : 0 ≤ y → int_to_nat Z N x < int_to_nat Z N y → x < y. Proof. intros. rewrite <-(int_to_nat_of_nonneg y) by easy. apply le_lt_trans with (f (int_to_nat Z N x)). now apply int_to_nat_le_r. now apply (strictly_order_preserving f). Qed. End contents. math-classes-8.19.0/theory/integers.v000066400000000000000000000126561460576051100175350ustar00rootroot00000000000000(* General results about arbitrary integer implementations. *) Require MathClasses.theory.nat_distance. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.naturals MathClasses.interfaces.abstract_algebra MathClasses.implementations.natpair_integers. Require Export MathClasses.interfaces.integers. (* Any two integer implementations are trivially isomorphic because of their initiality, but it's nice to have this stated in terms of integers_to_ring being self-inverse: *) Lemma to_ring_involutive `{Integers Z} Z2 `{Integers Z2} x : integers_to_ring Z2 Z (integers_to_ring Z Z2 x) = x. Proof. rapply (proj2 (@categories.initials_unique' _ _ _ _ _ _ (rings.object Z) (rings.object Z2) _ integers_initial _ integers_initial) tt x). Qed. Lemma to_ring_unique `{Integers Z} `{Ring R} (f: Z → R) {h: SemiRing_Morphism f} x : f x = integers_to_ring Z R x. Proof. symmetry. pose proof (rings.encode_morphism_and_ops (f:=f)). set (@varieties.arrow rings.theory _ _ _ (rings.encode_variety_and_ops _) _ _ _ (rings.encode_variety_and_ops _) (λ _, f) _). exact (integers_initial _ a tt x). Qed. Lemma to_ring_unique_alt `{Integers Z} `{Ring R} (f g: Z → R) `{!SemiRing_Morphism f} `{!SemiRing_Morphism g} x : f x = g x. Proof. now rewrite (to_ring_unique f), (to_ring_unique g). Qed. Lemma morphisms_involutive `{Integers Z} `{Integers Z2} (f: Z → Z2) (g: Z2 → Z) `{!SemiRing_Morphism f} `{!SemiRing_Morphism g} x : f (g x) = x. Proof. now apply (to_ring_unique_alt (f ∘ g) id). Qed. Lemma to_ring_twice `{Integers Z} `{Ring R1} `{Ring R2} (f : R1 → R2) (g : Z → R1) (h : Z → R2) `{!SemiRing_Morphism f} `{!SemiRing_Morphism g} `{!SemiRing_Morphism h} x : f (g x) = h x. Proof. now apply (to_ring_unique_alt (f ∘ g) h). Qed. Lemma to_ring_self `{Integers Z} (f : Z → Z) `{!SemiRing_Morphism f} x : f x = x. Proof. now apply (to_ring_unique_alt f id). Qed. (* A ring morphism from integers to another ring is injective if there's an injection in the other direction: *) Lemma to_ring_injective `{Integers Z} `{Ring R} (f: R → Z) (g: Z → R) `{!SemiRing_Morphism f} `{!SemiRing_Morphism g} : Injective g. Proof. repeat (split; try apply _). intros x y E. rewrite <-(to_ring_self (f ∘ g) x), <-(to_ring_self (f ∘ g) y). unfold compose. now rewrite E. Qed. #[global] Instance integers_to_integers_injective `{Integers Z} `{Integers Z2} (f: Z → Z2) `{!SemiRing_Morphism f} : Injective f. Proof. apply (to_ring_injective (integers_to_ring Z2 Z) _). Qed. #[global] Instance naturals_to_integers_injective `{Integers Z} `{Naturals N} (f: N → Z) `{!SemiRing_Morphism f} : Injective f. Proof. split; try apply _. intros x y E. apply (injective (cast N (SRpair N))). now rewrite <-2!(naturals.to_semiring_twice (integers_to_ring Z (SRpair N)) f (cast N (SRpair N))), E. Qed. Section retract_is_int. Context `{Integers Z} `{Ring Z2}. Context (f : Z → Z2) `{!Inverse f} `{!Surjective f} `{!SemiRing_Morphism f} `{!SemiRing_Morphism (f⁻¹)}. (* If we make this an instance, then instance resolution will often loop *) Definition retract_is_int_to_ring : IntegersToRing Z2 := λ Z2 _ _ _ _ _, integers_to_ring Z Z2 ∘ f⁻¹. Section for_another_ring. Context `{Ring R}. Instance: SemiRing_Morphism (integers_to_ring Z R ∘ f⁻¹) := {}. Context (h : Z2 → R) `{!SemiRing_Morphism h}. Lemma same_morphism: integers_to_ring Z R ∘ f⁻¹ = h. Proof with auto. intros x y F. rewrite <-F. transitivity ((h ∘ (f ∘ f⁻¹)) x). symmetry. unfold compose. apply (to_ring_unique (h ∘ f)). unfold compose. now rewrite jections.surjective_applied. Qed. End for_another_ring. (* If we make this an instance, then instance resolution will often loop *) Program Instance retract_is_int: Integers Z2 (U:=retract_is_int_to_ring). Next Obligation. unfold integers_to_ring, retract_is_int_to_ring. apply _. Qed. Next Obligation. apply integer_initial. intros. now apply same_morphism. Qed. End retract_is_int. Section contents. Context `{Integers Z}. Global Program Instance: ∀ x y: Z, Decision (x = y) | 10 := λ x y, match decide (integers_to_ring Z (SRpair nat) x = integers_to_ring Z (SRpair nat) y) with | left E => left _ | right E => right _ end. Next Obligation. now apply (injective (integers_to_ring Z (SRpair nat))). Qed. Global Program Instance slow_int_abs `{Naturals N} : IntAbs Z N | 10 := λ x, match int_abs_sig (SRpair N) N (integers_to_ring Z (SRpair N) x) with | inl (n↾E) => inl n | inr (n↾E) => inr n end. Next Obligation. apply (injective (integers_to_ring Z (SRpair N))). rewrite <-E. apply (naturals.to_semiring_twice _ _ _). Qed. Next Obligation. apply (injective (integers_to_ring Z (SRpair N))). rewrite rings.preserves_negate, <-E. now apply (naturals.to_semiring_twice _ _ _). Qed. Instance: PropHolds ((1:Z) ≠ 0). Proof. intros E. apply (rings.is_ne_0 1). apply (injective (naturals_to_semiring nat Z)). now rewrite rings.preserves_0, rings.preserves_1. Qed. Global Instance: ZeroProduct Z. Proof. intros x y E. destruct (zero_product (integers_to_ring Z (SRpair nat) x) (integers_to_ring Z (SRpair nat) y)). now rewrite <-rings.preserves_mult, E, rings.preserves_0. left. apply (injective (integers_to_ring Z (SRpair nat))). now rewrite rings.preserves_0. right. apply (injective (integers_to_ring Z (SRpair nat))). now rewrite rings.preserves_0. Qed. Global Instance: IntegralDomain Z := {}. End contents. math-classes-8.19.0/theory/jections.v000066400000000000000000000141571460576051100175310ustar00rootroot00000000000000Require Import MathClasses.theory.setoids MathClasses.interfaces.abstract_algebra. Local Existing Instance injective_mor. Local Existing Instance surjective_mor. Lemma injective_compose_cancel `{Equiv A} `{Equiv B} `{Equiv C} (f : B → C) `{!Injective f} `{!Setoid_Morphism (g : A → B)} `{!Setoid_Morphism (h : A → B)} : f ∘ g = f ∘ h → g = h. Proof. pose proof (setoidmor_a g). intros E. apply setoids.ext_equiv_applied_iff. intros x. apply (injective f). now apply E. Qed. Lemma surjective_applied `{Equiv A} `{Equiv B} (f : A → B) `{!Inverse f} `{!Surjective f} x : f (f⁻¹ x) = x. Proof. firstorder. Qed. #[global] Instance inverse_mor `{Bijective A B f} : Setoid_Morphism (f⁻¹). Proof. pose proof (setoidmor_a f). pose proof (setoidmor_b f). split; try apply _. intros x y E. apply (injective f). now rewrite !(surjective_applied f). Qed. Lemma bijective_cancel_left `{Equiv A} `{Equiv B} (f : A → B) `{!Inverse f} `{!Bijective f} x y : f x = y → x = f⁻¹ y. Proof. pose proof (setoidmor_b f). intros E. apply (injective f). now rewrite (surjective_applied f). Qed. Lemma bijective_cancel_inverse_left `{Equiv A} `{Equiv B} (f : A → B) `{!Inverse f} `{!Bijective f} x y : f⁻¹ x = y → x = f y. Proof. pose proof (setoidmor_a f). pose proof (setoidmor_b f). intros E. now rewrite <-E, (surjective_applied f). Qed. Lemma bijective_applied `{Equiv A} `{Equiv B} (f : A → B) `{!Inverse f} `{!Bijective f} x: f⁻¹ (f x) = x. Proof. pose proof (setoidmor_a f). pose proof (setoidmor_b f). symmetry. now apply (bijective_cancel_left f). Qed. Lemma bijective `{Equiv A} `{Equiv B} (f : A → B) `{!Inverse f} `{!Bijective f} : f⁻¹ ∘ f = id. (* a.k.a. "split-mono" *) Proof. pose proof (setoidmor_a f). apply ext_equiv_applied_iff, (bijective_applied f). Qed. Lemma injective_ne `{Equiv A} `{Equiv B} `(f : A → B) `{!Injective f} x y : x ≠ y → f x ≠ f y. Proof. intros E1 E2. apply E1. now apply (injective f). Qed. #[global] Instance id_inverse {A} : Inverse (@id A) := (@id A). Global Existing Instance id_morphism. #[global] Instance id_injective `{Setoid A} : Injective (@id A). Proof. split; try apply _. easy. Qed. #[global] Instance id_surjective `{Setoid A} : Surjective (@id A). Proof. split; try apply _. now repeat intro. Qed. #[global] Instance id_bijective `{Setoid A} : Bijective (@id A). Proof. split; try apply _. Qed. Section compositions. Context `{Equiv A} `{Equiv B} `{Equiv C} (g: A → B) (f: B → C) `{!Inverse f} `{!Inverse g}. Instance compose_inverse: Inverse (f ∘ g) := g⁻¹ ∘ f⁻¹. Instance compose_injective: Injective f → Injective g → Injective (f ∘ g). Proof. firstorder. Qed. Instance compose_surjective: Surjective f → Surjective g → Surjective (f ∘ g). Proof. split; try apply _. pose proof (setoidmor_b f). intros x y E. rewrite <-E. change (f (g (g⁻¹ (f⁻¹ x))) = x). now rewrite !surjective_applied. Qed. Instance compose_bijective: Bijective f → Bijective g → Bijective (f ∘ g) := {}. End compositions. #[global] Hint Extern 4 (Inverse (_ ∘ _)) => class_apply @compose_inverse : typeclass_instances. #[global] Hint Extern 4 (Injective (_ ∘ _)) => class_apply @compose_injective : typeclass_instances. #[global] Hint Extern 4 (Surjective (_ ∘ _)) => class_apply @compose_surjective : typeclass_instances. #[global] Hint Extern 4 (Bijective (_ ∘ _)) => class_apply @compose_bijective : typeclass_instances. Lemma alt_Build_Injective `{Equiv A} `{Equiv B} (f : A → B) `{!Inverse f} : Setoid_Morphism f → Setoid_Morphism (f⁻¹) → f⁻¹ ∘ f = id → Injective f. Proof. intros ?? E. pose proof (setoidmor_a f). pose proof (setoidmor_b f). split; try apply _. intros x y F. rewrite <-(ext_equiv_applied E x), <-(ext_equiv_applied E y). unfold compose. now rewrite F. Qed. Lemma alt_Build_Bijective `{Equiv A} `{Equiv B} (f : A → B) `{!Inverse f} : Setoid_Morphism f → Setoid_Morphism (f⁻¹) → f⁻¹ ∘ f = id → f ∘ f⁻¹ = id → Bijective f. Proof. intros. split. now apply (alt_Build_Injective f). split; auto. Qed. Definition inverse_inverse `{Inverse A B f} : Inverse (f⁻¹) := f. #[global] Hint Extern 4 (Inverse (_ ⁻¹)) => class_apply @inverse_inverse : typeclass_instances. Lemma flip_bijection `{Bijective A B f} : Bijective (f⁻¹). Proof. apply alt_Build_Bijective; try apply _. apply (surjective f). apply (bijective f). Qed. (* We use this instead of making flip_bijection a real instance, because otherwise it gets applied too eagerly, resulting in cycles. *) #[global] Hint Extern 4 (Bijective (_ ⁻¹)) => apply flip_bijection : typeclass_instances. Lemma inverse_involutive `(f : A → B) `{!Inverse f} : (f⁻¹)⁻¹ ≡ f. Proof. reflexivity. Qed. (* This second version is strictly for manual application. *) Lemma flip_bijection_back `{Equiv A} `{Equiv B} (f: A → B) `{!Inverse f} : Bijective (f⁻¹) → Bijective f. Proof. intro. apply (_: Bijective (f⁻¹⁻¹)). Qed. #[global] Instance injective_proper `{Equiv A} `{Equiv B} : Proper ((=) ==> (=)) (@Injective A B _ _). Proof. assert (∀ f g : A → B, f = g → Injective f → Injective g) as aux. intros f g E ?. pose proof (setoidmor_a f). pose proof (setoidmor_b f). split. intros x y ?. apply (injective f). now rewrite (ext_equiv_applied E x), (ext_equiv_applied E y). rewrite <-E; apply _. intros f g; split; intros; eapply aux; eauto. pose proof (setoidmor_a g). pose proof (setoidmor_b g). now symmetry. Qed. Lemma surjective_proper `{Equiv A} `{Equiv B} (f g : A → B) `{!Inverse f} `{!Inverse g} `{!Surjective g} : f = g → f⁻¹ = g⁻¹ → Surjective f. Proof. intros E1 E2. pose proof (setoidmor_a g). pose proof (setoidmor_b g). split. intros ? ? E3. change (f (f⁻¹ x) = y). rewrite <-E3, (ext_equiv_applied E1 _), (ext_equiv_applied E2 _). now apply surjective_applied. rewrite E1; apply _. Qed. Ltac setoid_inject := match goal with | E : _ = ?f _ |- _ => apply (injective f) in E | E : ?f _ = _ |- _ => apply (injective f) in E | E : _ ≡ _ |- ?G => change (id G); injection E; clear E; intros; unfold id at 1 end. math-classes-8.19.0/theory/lattices.v000066400000000000000000000161321460576051100175160ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.theory.groups. Require MathClasses.varieties.semigroups MathClasses.varieties.monoids. #[global] Instance bounded_sl_is_sl `{BoundedSemiLattice L} : SemiLattice L. Proof. repeat (split; try apply _). Qed. #[global] Instance bounded_join_sl_is_join_sl `{BoundedJoinSemiLattice L} : JoinSemiLattice L. Proof. repeat (split; try apply _). Qed. #[global] Instance bounded_sl_mor_is_sl_mor `{H : BoundedJoinSemiLattice_Morphism A B f} : JoinSemiLattice_Morphism f. Proof. destruct H. split; apply _. Qed. Lemma preserves_join `{JoinSemiLattice_Morphism L K f} x y : f (x ⊔ y) = f x ⊔ f y. Proof preserves_sg_op x y. Lemma preserves_bottom `{BoundedJoinSemiLattice_Morphism L K f} : f ⊥ = ⊥. Proof (preserves_mon_unit (f:=f)). Lemma preserves_meet `{MeetSemiLattice_Morphism L K f} x y : f (x ⊓ y) = f x ⊓ f y. Proof preserves_sg_op x y. Section bounded_join_sl_props. Context `{BoundedJoinSemiLattice L}. Instance join_bottom_l: LeftIdentity (⊔) ⊥ := _. Instance join_bottom_r: RightIdentity (⊔) ⊥ := _. End bounded_join_sl_props. Section lattice_props. Context `{Lattice L}. Definition meet_join_absorption x y : x ⊓ (x ⊔ y) = x := absorption x y. Definition join_meet_absorption x y : x ⊔ (x ⊓ y) = x := absorption x y. End lattice_props. Section distributive_lattice_props. Context `{DistributiveLattice L}. Instance join_meet_distr_l: LeftDistribute (⊔) (⊓). Proof join_meet_distr_l _. Global Instance join_meet_distr_r: RightDistribute (⊔) (⊓). Proof. intros x y z. rewrite !(commutativity _ z). now apply distribute_l. Qed. Global Instance meet_join_distr_l: LeftDistribute (⊓) (⊔). Proof. intros x y z. rewrite distribute_l. rewrite distribute_r. rewrite (idempotency (⊔) x). rewrite (commutativity y x), meet_join_absorption. rewrite <-(meet_join_absorption x z) at 1. rewrite <-associativity. now rewrite <-distribute_r. Qed. Global Instance meet_join_distr_r: RightDistribute (⊓) (⊔). Proof. intros x y z. rewrite !(commutativity _ z). now apply distribute_l. Qed. Lemma distribute_alt x y z : (x ⊓ y) ⊔ (x ⊓ z) ⊔ (y ⊓ z) = (x ⊔ y) ⊓ (x ⊔ z) ⊓ (y ⊔ z). Proof. rewrite (distribute_r x y (x ⊓ z)), join_meet_absorption. rewrite (distribute_r _ _ (y ⊓ z)). rewrite (distribute_l x y z). rewrite (commutativity y (x ⊓ z)), <-(associativity _ y). rewrite join_meet_absorption. rewrite (distribute_r x z y). rewrite (commutativity z y). rewrite (commutativity (x ⊔ y) (x ⊔ z)). rewrite associativity, <-(associativity (x ⊔ z)). rewrite (idempotency _ _). now rewrite (commutativity (x ⊔ z) (x ⊔ y)). Qed. End distributive_lattice_props. Section lower_bounded_lattice. Context `{Lattice L} `{Bottom L} `{!BoundedJoinSemiLattice L}. Global Instance meet_bottom_l: LeftAbsorb (⊓) ⊥. Proof. intros x. now rewrite <-(join_bottom_l x), absorption. Qed. Global Instance meet_bottom_r: RightAbsorb (⊓) ⊥. Proof. intros x. now rewrite commutativity, left_absorb. Qed. End lower_bounded_lattice. Section from_another_sl. Context `{SemiLattice A} `{Setoid B} `{Bop : SgOp B} (f : B → A) `{!Injective f} (op_correct : ∀ x y, f (x & y) = f x & f y). Lemma projected_sl: SemiLattice B. Proof. split. now apply (projected_com_sg f). repeat intro; apply (injective f). now rewrite !op_correct, (idempotency (&) _). Qed. End from_another_sl. Section from_another_bounded_sl. Context `{BoundedSemiLattice A} `{Setoid B} `{Bop : SgOp B} `{Bunit : MonUnit B} (f : B → A) `{!Injective f} (op_correct : ∀ x y, f (x & y) = f x & f y) (unit_correct : f mon_unit = mon_unit). Lemma projected_bounded_sl: BoundedSemiLattice B. Proof. split. now apply (projected_com_monoid f). repeat intro; apply (injective f). now rewrite op_correct, (idempotency (&) _). Qed. End from_another_bounded_sl. #[global] Instance id_join_sl_morphism `{JoinSemiLattice A} : JoinSemiLattice_Morphism (@id A). Proof. firstorder. Qed. #[global] Instance id_meet_sl_morphism `{MeetSemiLattice A} : MeetSemiLattice_Morphism (@id A). Proof. firstorder. Qed. #[global] Instance id_bounded_join_sl_morphism `{BoundedJoinSemiLattice A} : BoundedJoinSemiLattice_Morphism (@id A). Proof. firstorder. Qed. #[global] Instance id_lattice_morphism `{Lattice A} : Lattice_Morphism (@id A). Proof. firstorder. Qed. Section morphism_composition. Context `{Equiv A} `{Join A} `{Meet A} `{Bottom A} `{Equiv B}`{Join B} `{Meet B} `{Bottom B} `{Equiv C}`{Join C} `{Meet C} `{Bottom C} (f : A → B) (g : B → C). Instance compose_join_sl_morphism: JoinSemiLattice_Morphism f → JoinSemiLattice_Morphism g → JoinSemiLattice_Morphism (g ∘ f). Proof. split; try apply _; firstorder. Qed. Instance compose_meet_sl_morphism: MeetSemiLattice_Morphism f → MeetSemiLattice_Morphism g → MeetSemiLattice_Morphism (g ∘ f). Proof. split; try apply _; firstorder. Qed. Instance compose_bounded_join_sl_morphism: BoundedJoinSemiLattice_Morphism f → BoundedJoinSemiLattice_Morphism g → BoundedJoinSemiLattice_Morphism (g ∘ f). Proof. split; try apply _; firstorder. Qed. Instance compose_lattice_morphism: Lattice_Morphism f → Lattice_Morphism g → Lattice_Morphism (g ∘ f). Proof. split; try apply _; firstorder. Qed. Instance invert_join_sl_morphism: ∀ `{!Inverse f}, Bijective f → JoinSemiLattice_Morphism f → JoinSemiLattice_Morphism (f⁻¹). Proof. split; try apply _; firstorder. Qed. Instance invert_meet_sl_morphism: ∀ `{!Inverse f}, Bijective f → MeetSemiLattice_Morphism f → MeetSemiLattice_Morphism (f⁻¹). Proof. split; try apply _; firstorder. Qed. Instance invert_bounded_join_sl_morphism: ∀ `{!Inverse f}, Bijective f → BoundedJoinSemiLattice_Morphism f → BoundedJoinSemiLattice_Morphism (f⁻¹). Proof. split; try apply _; firstorder. Qed. Instance invert_lattice_morphism: ∀ `{!Inverse f}, Bijective f → Lattice_Morphism f → Lattice_Morphism (f⁻¹). Proof. split; try apply _; firstorder. Qed. End morphism_composition. #[global] Hint Extern 4 (JoinSemiLattice_Morphism (_ ∘ _)) => class_apply @compose_join_sl_morphism : typeclass_instances. #[global] Hint Extern 4 (MeetSemiLattice_Morphism (_ ∘ _)) => class_apply @compose_meet_sl_morphism : typeclass_instances. #[global] Hint Extern 4 (BoundedJoinSemiLattice_Morphism (_ ∘ _)) => class_apply @compose_bounded_join_sl_morphism : typeclass_instances. #[global] Hint Extern 4 (Lattice_Morphism (_ ∘ _)) => class_apply @compose_lattice_morphism : typeclass_instances. #[global] Hint Extern 4 (JoinSemiLattice_Morphism (_⁻¹)) => class_apply @invert_join_sl_morphism : typeclass_instances. #[global] Hint Extern 4 (MeetSemiLattice_Morphism (_⁻¹)) => class_apply @invert_meet_sl_morphism : typeclass_instances. #[global] Hint Extern 4 (BoundedJoinSemiLattice_Morphism (_⁻¹)) => class_apply @invert_bounded_join_sl_morphism : typeclass_instances. #[global] Hint Extern 4 (Lattice_Morphism (_⁻¹)) => class_apply @invert_lattice_morphism : typeclass_instances. math-classes-8.19.0/theory/monads.v000066400000000000000000000235011460576051100171650ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.monads MathClasses.theory.functors. #[global] Instance default_mon_join `{MonadBind M} : MonadJoin M | 20 := λ _, bind id. #[global] Instance default_mon_map `{MonadReturn M} `{MonadBind M} : SFmap M | 20 := λ _ _ f, bind (ret ∘ f). #[global] Instance default_mon_bind `{SFmap M} `{MonadJoin M} : MonadBind M | 20 := λ _ _ f, join ∘ (sfmap f). #[global] Hint Extern 0 (ProperProxy (@respectful _ _ _ _) _) => class_apply @proper_proper_proxy : typeclass_instances. #[global] Instance equiv_ext_equiv `{Equiv A} `{Equiv B} : Setoid A -> Setoid B -> Proper ((equiv ==> equiv) ==> (equiv ==> equiv) ==> flip impl) (@equiv _ (@ext_equiv A _ B _)). Proof. unfold ext_equiv. repeat (red; intros). assert ((equiv ==> equiv)%signature x x0). eapply transitivity. eauto. eapply transitivity. eauto. eapply symmetry. eauto. eapply H7. eapply H6. Qed. #[global] Instance equiv_ext_equiv_partial `{Equiv A} `{Equiv B} (f : A -> B) : Setoid A -> Setoid B -> Proper (equiv ==> equiv) f -> Proper ((equiv ==> equiv) ==> flip impl) (@equiv _ (@ext_equiv A _ B _) f). Proof. intros. partial_application_tactic; eauto. apply equiv_ext_equiv; eauto. Qed. Section monad. Context `{Monad M}. Lemma bind_lunit_applied `{Equiv A} `{Setoid B} `{!Setoid_Morphism (f : A → M B)} (x : A) : ret x ≫= f = f x. Proof. pose proof (setoidmor_a f). now apply bind_lunit. Qed. Lemma bind_runit_applied `{Setoid A} (m : M A) : m ≫= ret = m. Proof. now apply bind_runit. Qed. Lemma bind_assoc_applied `{Equiv A} `{Equiv B} `{Setoid C} `{!Setoid_Morphism (f : A → M B)} `{!Setoid_Morphism (g : B → M C)} (m : M A) : (m ≫= f) ≫= g = x ← m ; f x ≫= g. Proof. pose proof (setoidmor_a f). now apply bind_assoc. Qed. Global Instance ret_mor `{Setoid A} : Setoid_Morphism (@ret _ _ A) := {}. Global Instance bind_mor `{Equiv A} `{Setoid B} `{!Setoid_Morphism (f : A → M B)} : Setoid_Morphism (bind f). Proof. pose proof (setoidmor_a f). split; try apply _. Qed. Definition liftM2 `(f: A → B → C) (m : M A) (n : M B) : M C := x ← m ; y ← n ; ret (f x y). Section to_strong_monad. Context `{MonadJoin M} `{SFmap M} (map_proper : ∀ `{Setoid A} `{Setoid B}, Proper (((=) ==> (=)) ==> ((=) ==> (=))) (@sfmap M _ A B)) (map_correct : ∀ `{Equiv A} `{Equiv B} `{!Setoid_Morphism (f : A → B)}, sfmap f = bind (ret ∘ f)) (join_correct : ∀ `{Setoid A}, join = bind id). Existing Instance map_proper. Let bind_correct `{Equiv A} `{Setoid B} `{!Setoid_Morphism (f : A → M B)} : bind f = join ∘ sfmap f. Proof. pose proof (setoidmor_a f). pose proof (setoidmor_b f). rewrite join_correct, map_correct by apply _. rewrite bind_assoc. change (bind f = bind ((bind id ∘ ret) ∘ f)). rewrite bind_lunit. now apply setoids.ext_equiv_refl. Qed. Instance: SFunctor M. Proof. split; try apply _. intros A ? ?. rewrite map_correct by apply _. now apply bind_runit. intros A ? B ? C ? f ? g ?. pose proof (setoidmor_a g). pose proof (setoidmor_b g). pose proof (setoidmor_b f). rewrite !map_correct by apply _. rewrite bind_assoc. change (bind (ret ∘ (f ∘ g)) = bind ((bind (ret ∘ f) ∘ ret) ∘ g)). rewrite bind_lunit. now apply setoids.ext_equiv_refl. Qed. Instance: ∀ `{Setoid A}, Setoid_Morphism (@join _ _ A). Proof. split; try apply _. intros x y E1. assert (∀ z, join z = bind id z) as E2 by (intros; now apply join_correct). now rewrite !E2, E1. Qed. Instance monad_strong_monad: StrongMonad M. Proof. split; try apply _. intros A ? B ? f ?. pose proof (setoidmor_a f). pose proof (setoidmor_b f). rewrite map_correct by apply _. rewrite bind_lunit. now apply setoids.ext_equiv_refl. intros A ? B ? f ?. pose proof (setoidmor_a f). pose proof (setoidmor_b f). rewrite <-bind_correct. rewrite !join_correct by apply _. rewrite map_correct by apply _. rewrite bind_assoc. now apply setoids.ext_equiv_refl. intros A ??. rewrite join_correct by apply _. rewrite bind_lunit. now apply setoids.ext_equiv_refl. intros A ??. rewrite <-bind_correct. rewrite bind_runit. now apply setoids.ext_equiv_refl. intros A ??. rewrite <-bind_correct. rewrite !join_correct by apply _. rewrite bind_assoc. now apply setoids.ext_equiv_refl. Qed. Instance monad_full_monad: FullMonad M. Proof. split; try apply _; auto. Qed. End to_strong_monad. Instance monad_default_full_monad: FullMonad M. Proof. apply monad_full_monad; unfold sfmap, default_mon_map. intros A ?? B ?? f g E1 m n E2. apply mon_bind_proper. intros x y E3. now apply mon_ret_proper, E1. easy. intros A ? B ? f ??? E. pose proof (setoidmor_a f). pose proof (setoidmor_b f). now rewrite E. intros A ?? ?? E. unfold join, default_mon_join. now rewrite E. Qed. End monad. Section strong_monad. Context `{StrongMonad M}. Global Instance sret_mor `{Setoid A} : Setoid_Morphism (@ret _ _ A) := {}. Global Instance join_mor `{Setoid A} : Setoid_Morphism (@join _ _ A) := {}. Hint Immediate setoidmor_a : typeclass_instances. Hint Immediate setoidmor_b : typeclass_instances. Lemma sfmap_ret_applied `{Equiv A} `{Equiv B} `{!Setoid_Morphism (f : A → B)} (x : A) : sfmap f (ret x) = ret (f x). Proof. now apply sfmap_ret. Qed. Lemma sfmap_join_applied `{Equiv A} `{Equiv B} `{!Setoid_Morphism (f : A → B)} (m : M (M A)) : sfmap f (join m) = join (sfmap (sfmap f) m). Proof. now apply sfmap_join. Qed. Lemma join_ret_applied `{Setoid A} (m : M A) : join (ret m) = m. Proof. now apply join_ret. Qed. Lemma join_sfmap_ret_applied `{Setoid A} (m : M A): join (sfmap ret m) = m. Proof. now apply join_sfmap_ret. Qed. Lemma join_sfmap_join_applied `{Setoid A} (m : M (M (M A))) : join (sfmap join m) = join (join m). Proof. now apply join_sfmap_join. Qed. Section to_monad. Context `{MonadBind M} (bind_proper : ∀ `{Setoid A} `{Setoid B}, Proper (((=) ==> (=)) ==> ((=) ==> (=))) (@bind M _ A B)) (bind_correct : ∀ `{Equiv A} `{Setoid B} `{!Setoid_Morphism (f : A → M B)}, bind f = join ∘ sfmap f). Instance: ∀ `{Equiv A} `{Setoid B} `{!Setoid_Morphism (f : A → M B)}, Setoid_Morphism (bind f). Proof. intros. split; try apply _. Qed. Let bind_correct_applied `{Equiv A} `{Setoid B} `{!Setoid_Morphism (f : A → M B)} m : bind f m = join (sfmap f m). Proof. now eapply bind_correct. Qed. Instance strong_monad_monad: Monad M. Proof. split; try apply _. intros A ? B ?? f ?. pose proof (setoidmor_a f). pose proof (setoidmor_b f). rewrite bind_correct by apply _. rewrite compose_assoc, sfmap_ret. rewrite <-compose_assoc, join_ret. now apply setoids.ext_equiv_refl. intros A ? ?. rewrite bind_correct by apply _. now apply join_sfmap_ret. intros A ? B ? C ?? f ? g ? m n E. pose proof (setoidmor_a f). pose proof (setoidmor_a g). unfold compose at 1. rewrite !bind_correct_applied. rewrite bind_correct by apply _. rewrite sfmap_join_applied. rewrite !sfmap_comp_applied. rewrite join_sfmap_join_applied. now rewrite E. Qed. Instance strong_monad_full_monad: FullMonad M. Proof. split; try apply _; auto. Qed. End to_monad. Instance strong_monad_default_full_monad: FullMonad M. Proof. apply strong_monad_full_monad; unfold bind, default_mon_bind. intros A ?? B ?? f g E1 m n E2. apply smon_join_proper. apply sfmap_proper; intuition. intros A ? B ?? f ? ?? E. now rewrite E. Qed. End strong_monad. Section full_monad. Context `{FullMonad M}. Lemma bind_as_join_sfmap_applied `{Equiv A} `{Setoid B} `{!Setoid_Morphism (f : A → M B)} (m : M A) : m ≫= f = join (sfmap f m). Proof. pose proof (setoidmor_a f). now apply bind_as_join_sfmap. Qed. Lemma sfmap_as_bind_ret `{Equiv A} `{Equiv B} `{!Setoid_Morphism (f : A → B)} : sfmap f = bind (ret ∘ f). Proof. pose proof (setoidmor_a f). pose proof (setoidmor_b f). rewrite bind_as_join_sfmap. rewrite sfmap_comp. rewrite <-compose_assoc. rewrite join_sfmap_ret. now apply setoids.ext_equiv_refl. Qed. Lemma sfmap_as_bind_ret_applied `{Equiv A} `{Equiv B} `{!Setoid_Morphism (f : A → B)} (m : M A) : sfmap f m = x ← m ; ret (f x). Proof. pose proof (setoidmor_a f). now apply sfmap_as_bind_ret. Qed. Lemma join_as_bind `{Setoid A} : join = bind id. Proof. rewrite bind_as_join_sfmap. rewrite sfmap_id. now apply setoids.ext_equiv_refl. Qed. Lemma join_as_bind_applied `{Setoid A} (m : M (M A)) : join m = m ≫= id. Proof. now apply join_as_bind. Qed. Lemma join_spec_applied_alt `{Setoid A} (m : M (M A)) : join m = x ← m ; x. Proof. now apply join_as_bind. Qed. Lemma bind_twice `{Equiv A} `{Equiv B} `{Setoid C} `{!Setoid_Morphism (f : B → M C)} `{!Setoid_Morphism (g : A → M B)} : bind (bind f) = bind f ∘ join. Proof. pose proof (setoidmor_a f). pose proof (setoidmor_b f). pose proof (setoidmor_b g). rewrite join_as_bind. rewrite bind_assoc. now apply setoids.ext_equiv_refl. Qed. Lemma bind_twice_applied `{Equiv A} `{Equiv B} `{Setoid C} `{!Setoid_Morphism (f : B → M C)} `{!Setoid_Morphism (g : A → M B)} (m : M (M B)) : m ≫= bind f = join m ≫= f. Proof. pose proof (setoidmor_a f). now apply bind_twice. Qed. Lemma bind_join `{Setoid A} : bind join = join ∘ join. Proof. rewrite !join_as_bind. rewrite bind_assoc. now apply setoids.ext_equiv_refl. Qed. Lemma bind_join_applied `{Setoid A} (m : M (M (M A))) : m ≫= join = join (join m). Proof. now apply bind_join. Qed. End full_monad. math-classes-8.19.0/theory/monoid_normalization.v000066400000000000000000000075661460576051100221540ustar00rootroot00000000000000From Coq Require Import Lia. Require Import MathClasses.interfaces.abstract_algebra MathClasses.theory.ua_packed. Require MathClasses.interfaces.universal_algebra MathClasses.varieties.monoids. Notation msig := varieties.monoids.sig. Notation Op := (universal_algebra.Op msig False). Notation App := (universal_algebra.App msig False _ _). Import universal_algebra. Section contents. (* Ideally, we would like to operate exclusively on the universal term representation(s). If Coq had decent support for dependent pattern matching, this would be no problem. Unfortunately, Coq does not, and so we resort to defining an ad-hoc data type for monoidal expressions, with nasty conversions to and from the universal term representation(s): *) Context (V: Type). Notation uaTerm := (universal_algebra.Term0 msig V tt). Notation Applied := (@ua_packed.Applied msig V tt). Inductive Term := Var (v: V) | Unit | Comp (x y: Term). Fixpoint to_ua (e: Term): Applied := match e with | Var v => ua_packed.AppliedVar msig v tt | Unit => ua_packed.AppliedOp msig monoids.one (ua_packed.NoMoreArguments msig tt) | Comp x y => ua_packed.AppliedOp msig monoids.mult (MoreArguments msig tt _ (to_ua x) (MoreArguments msig tt _ (to_ua y) (NoMoreArguments msig tt))) end. Definition from_ua (t: Applied): { r: Term | to_ua r ≡ t }. Proof with reflexivity. refine (better_Applied_rect msig (λ s, match s return (ua_packed.Applied msig s → Type) with | tt => λ t, {r : Term | to_ua r ≡ t} end) _ _ t). simpl. intros []; simpl. intros. exists (Comp (` (fst (forallSplit msig _ _ X))) (` ((fst (forallSplit msig _ _ (snd (forallSplit msig _ _ X))))))). do 3 dependent destruction x. simpl in *. destruct X. destruct p, s. destruct s0. simpl. subst... exists Unit. dependent destruction x... intros v []. exists (Var v)... Defined. Fixpoint measure (e: Term): nat := match e with | Var v => 0%nat | Unit => 1%nat | Comp x y => S (2 * measure x + measure y) end. Context `{Monoid M}. Notation eval vs := (universal_algebra.eval msig (λ _, (vs: V → M))). Program Fixpoint internal_simplify (t: Term) {measure (measure t)}: { r: Term | ∀ v, eval v (curry (to_ua r)) = eval v (curry (to_ua t)) } := match t with | Var _ => t | Unit => t | Comp Unit y => internal_simplify y | Comp x Unit => internal_simplify x | Comp ((Var _) as x) y => Comp x (internal_simplify y) | Comp (Comp x y) z => internal_simplify (Comp x (Comp y z)) end. Solve Obligations with (program_simpl; simpl; try apply reflexivity; clear internal_simplify; lia). Next Obligation. destruct internal_simplify. simpl. rewrite e. transitivity (mon_unit & universal_algebra.eval msig (λ _, v) (curry (to_ua y))). symmetry. apply left_identity. reflexivity. Qed. Next Obligation. destruct internal_simplify. simpl. rewrite e. transitivity (universal_algebra.eval msig (λ _, v) (curry (to_ua x)) & mon_unit). symmetry. apply right_identity. reflexivity. Qed. Next Obligation. destruct internal_simplify. simpl. rewrite e. reflexivity. Qed. Next Obligation. destruct internal_simplify. simpl. rewrite e. simpl. apply associativity. Qed. Program Definition simplify (t: uaTerm): { r: uaTerm | ∀ v, eval v r = eval v t } := curry (to_ua (internal_simplify (from_ua (decode0 t)))). Next Obligation. destruct @internal_simplify. simpl. destruct @from_ua in e. simpl in *. rewrite e. rewrite e0. rewrite @curry_decode0. reflexivity. Qed. (* This would be a nice theorem to prove: Require varieties.open_terms. Instance: Equiv V := eq. Goal ∀ (x y: uaTerm), open_terms.ee msig msig monoids.Laws (ne_list.one tt) x y → ` (simplify x) ≡ ` (simplify y). Proof. *) End contents. math-classes-8.19.0/theory/nat_distance.v000066400000000000000000000054531460576051100203460ustar00rootroot00000000000000Require MathClasses.orders.naturals MathClasses.implementations.peano_naturals. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals MathClasses.interfaces.orders MathClasses.interfaces.additional_operations. Section contents. Context `{Naturals N}. Add Ring N : (rings.stdlib_semiring_theory N). (* NatDistance instances are all equivalent, because their behavior is fully determined by the specification. *) Lemma nat_distance_unique_respectful {a b : NatDistance N} : ((=) ==> (=) ==> (=))%signature (nat_distance (nd:=a)) (nat_distance (nd:= b)). Proof. intros x1 y1 E x2 y2 F. unfold nat_distance, nat_distance_sig. destruct a as [[z1 A]|[z1 A]], b as [[z2 B]|[z2 B]]; simpl. apply (left_cancellation (+) x1). now rewrite A, E, B. destruct (naturals.zero_sum z1 z2). apply (left_cancellation (+) x1). rewrite associativity, A, F, B, E. ring. transitivity 0; intuition. destruct (naturals.zero_sum z1 z2). rewrite commutativity. apply (left_cancellation (+) y1). rewrite associativity, B, <-F, A, E. ring. transitivity 0; intuition. apply (left_cancellation (+) x2). now rewrite A, E, F, B. Qed. Lemma nat_distance_unique {a b: NatDistance N} {x y : N} : nat_distance (nd:=a) x y = nat_distance (nd:=b) x y. Proof. now apply nat_distance_unique_respectful. Qed. Context `{!NatDistance N}. Global Instance nat_distance_proper : Proper ((=) ==> (=) ==> (=)) nat_distance. Proof. apply nat_distance_unique_respectful. Qed. End contents. (* An existing instance of [CutMinus] allows to create an instance of [NatDistance] *) #[global] Program Instance natdistance_cut_minus `{Naturals N} `{Apart N} `{!TrivialApart N} `{!FullPseudoSemiRingOrder Nle Nlt} `{!CutMinusSpec N cm} `{∀ x y, Decision (x ≤ y)} : NatDistance N := λ x y, if decide_rel (≤) x y then inl (y ∸ x) else inr (x ∸ y). Next Obligation. rewrite commutativity. now apply cut_minus_le. Qed. Next Obligation. rewrite commutativity. now apply cut_minus_le, orders.le_flip. Qed. (* Using the preceding instance we can make an instance for arbitrary models of the naturals by translation into [nat] on which we already have a [CutMinus] instance. *) Global Program Instance natdistance_default `{Naturals N} : NatDistance N | 10 := λ x y, match nat_distance_sig (naturals_to_semiring N nat x) (naturals_to_semiring N nat y) with | inl (n↾E) => inl (naturals_to_semiring nat N n) | inr (n↾E) => inr (naturals_to_semiring nat N n) end. Next Obligation. rewrite <-(naturals.to_semiring_involutive N nat y), <-E. now rewrite rings.preserves_plus, (naturals.to_semiring_involutive _ _). Qed. Next Obligation. rewrite <-(naturals.to_semiring_involutive N nat x), <-E. now rewrite rings.preserves_plus, (naturals.to_semiring_involutive _ _). Qed. math-classes-8.19.0/theory/nat_pow.v000066400000000000000000000145651460576051100173650ustar00rootroot00000000000000Require MathClasses.theory.naturals. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals MathClasses.interfaces.orders MathClasses.interfaces.additional_operations. (* * Properties of Nat Pow *) Section nat_pow_properties. Context `{SemiRing A} `{Naturals B} `{!NatPowSpec A B pw}. Add Ring A: (rings.stdlib_semiring_theory A). Add Ring B: (rings.stdlib_semiring_theory B). Global Instance: Proper ((=) ==> (=) ==> (=)) (^) | 0. Proof nat_pow_proper. Global Instance nat_pow_mor_1: ∀ x : A, Setoid_Morphism (x^) | 0. Proof. split; try apply _. Qed. Global Instance nat_pow_mor_2: ∀ n : B, Setoid_Morphism (^n) | 0. Proof. split; try apply _. solve_proper. Qed. Lemma nat_pow_base_0 (n : B) : n ≠ 0 → 0 ^ n = 0. Proof. pattern n. apply naturals.induction; clear n. solve_proper. intros E. now destruct E. intros. rewrite nat_pow_S. ring. Qed. Global Instance nat_pow_1: RightIdentity (^) (1:B). Proof. intro. assert ((1:B) = 1 + 0) as E by ring. rewrite E. rewrite nat_pow_S, nat_pow_0. ring. Qed. Lemma nat_pow_2 x : x ^ (2:B) = x * x. Proof. now rewrite nat_pow_S, nat_pow_1. Qed. Lemma nat_pow_3 x : x ^ (3:B) = x * (x * x). Proof. now rewrite nat_pow_S, nat_pow_2. Qed. Lemma nat_pow_4 x : x ^ (4:B) = x * (x * (x * x)). Proof. now rewrite nat_pow_S, nat_pow_3. Qed. Global Instance nat_pow_base_1: LeftAbsorb (^) 1. Proof. intro. pattern y. apply naturals.induction; clear y. solve_proper. apply nat_pow_0. intros n E. rewrite nat_pow_S. rewrite E. ring. Qed. Lemma nat_pow_exp_plus (x : A) (n m : B) : x ^ (n + m) = x ^ n * x ^ m. Proof. pattern n. apply naturals.induction; clear n. solve_proper. rewrite nat_pow_0, left_identity. ring. intros n E. rewrite <-associativity. rewrite 2!nat_pow_S. rewrite E. ring. Qed. Lemma nat_pow_base_mult (x y : A) (n : B) : (x * y) ^ n = x ^ n * y ^ n. Proof. pattern n. apply naturals.induction; clear n. solve_proper. rewrite ?nat_pow_0. ring. intros n E. rewrite ?nat_pow_S. rewrite E. ring. Qed. Lemma nat_pow_exp_mult (x : A) (n m : B) : x ^ (n * m) = (x ^ n) ^ m. Proof. pattern m. apply naturals.induction; clear m. solve_proper. rewrite right_absorb. now rewrite ?nat_pow_0. intros m E. rewrite nat_pow_S, <-E. rewrite distribute_l, right_identity. now rewrite nat_pow_exp_plus. Qed. Instance nat_pow_ne_0 `{!NoZeroDivisors A} `{!PropHolds ((1:A) ≠ 0)} (x : A) (n : B) : PropHolds (x ≠ 0) → PropHolds (x ^ n ≠ 0). Proof. pattern n. apply naturals.induction; clear n. solve_proper. intros. rewrite nat_pow_0. now apply (rings.is_ne_0 1). intros n E F G. rewrite nat_pow_S in G. unfold PropHolds in *. apply (no_zero_divisors x); split; eauto. Qed. Context `{Apart A} `{!FullPseudoSemiRingOrder (A:=A) Ale Alt} `{PropHolds (1 ≶ 0)}. Instance: StrongSetoid A := pseudo_order_setoid. Instance nat_pow_apart_0 (x : A) (n : B) : PropHolds (x ≶ 0) → PropHolds (x ^ n ≶ 0). Proof. pattern n. apply naturals.induction; clear n. solve_proper. intros. now rewrite nat_pow_0. intros n E F. rewrite nat_pow_S. rewrite <-(rings.mult_0_r x). apply (strong_left_cancellation (.*.) x). now apply E. Qed. Instance nat_pow_nonneg (x : A) (n : B) : PropHolds (0 ≤ x) → PropHolds (0 ≤ x ^ n). Proof. intros. pattern n. apply naturals.induction; clear n. solve_proper. rewrite nat_pow_0. apply _. intros. rewrite nat_pow_S. apply _. Qed. Instance nat_pow_pos (x : A) (n : B) : PropHolds (0 < x) → PropHolds (0 < x ^ n). Proof. rewrite !lt_iff_le_apart. intros [? ?]. split. now apply nat_pow_nonneg. symmetry. apply nat_pow_apart_0. red. now symmetry. Qed. Lemma nat_pow_ge_1 (x : A) (n : B) : 1 ≤ x → 1 ≤ x ^ n. Proof. intros. pattern n. apply naturals.induction. solve_proper. now rewrite nat_pow_0. intros. rewrite nat_pow_S. now apply semirings.ge_1_mult_compat. Qed. End nat_pow_properties. (* Due to bug #2528 *) #[global] Hint Extern 18 (PropHolds (_ ^ _ ≠ 0)) => eapply @nat_pow_ne_0 : typeclass_instances. #[global] Hint Extern 18 (PropHolds (_ ^ _ ≶ 0)) => eapply @nat_pow_apart_0 : typeclass_instances. #[global] Hint Extern 18 (PropHolds (0 ≤ _ ^ _)) => eapply @nat_pow_nonneg : typeclass_instances. #[global] Hint Extern 18 (PropHolds (0 < _ ^ _)) => eapply @nat_pow_pos : typeclass_instances. Section preservation. Context `{Naturals B} `{SemiRing A1} `{!NatPowSpec A1 B pw1} `{SemiRing A2} `{!NatPowSpec A2 B pw2} {f : A1 → A2} `{!SemiRing_Morphism f}. Add Ring B2 : (rings.stdlib_semiring_theory B). Lemma preserves_nat_pow x (n : B) : f (x ^ n) = (f x) ^ n. Proof. revert n. apply naturals.induction. solve_proper. rewrite nat_pow_0, nat_pow_0. now apply rings.preserves_1. intros n E. rewrite nat_pow_S, rings.preserves_mult, E. now rewrite nat_pow_S. Qed. End preservation. Section exp_preservation. Context `{SemiRing A} `{Naturals B1} `{Naturals B2} `{!NatPowSpec A B1 pw1} `{!NatPowSpec A B2 pw2} {f : B1 → B2} `{!SemiRing_Morphism f}. Lemma preserves_nat_pow_exp x (n : B1) : x ^ (f n) = x ^ n. Proof. revert n. apply naturals.induction. solve_proper. rewrite rings.preserves_0. now rewrite 2!nat_pow_0. intros n E. rewrite rings.preserves_plus, rings.preserves_1. rewrite 2!nat_pow_S. now rewrite E. Qed. End exp_preservation. (* Very slow default implementation by translation into Peano *) Section nat_pow_default. Context `{SemiRing A}. Global Instance nat_pow_peano: Pow A nat := fix nat_pow_rec (x: A) (n : nat) : A := match n with | 0 => 1 | S n => x * @pow _ _ nat_pow_rec x n end. Instance: Proper ((=) ==> (=) ==> (=)) nat_pow_peano. Proof. intros ? ? E a ? []. induction a; try easy. simpl. now rewrite IHa, E. Qed. Global Instance: NatPowSpec A nat nat_pow_peano. Proof. split; try apply _; easy. Qed. Context `{Naturals B}. Global Instance default_nat_pow: Pow A B | 10 := λ x n, x ^ naturals_to_semiring B nat n. Global Instance: NatPowSpec A B default_nat_pow. Proof. split; unfold pow, default_nat_pow. solve_proper. intros x. now rewrite rings.preserves_0. intros x n. now rewrite rings.preserves_plus, rings.preserves_1. Qed. End nat_pow_default. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Opaque default_nat_pow. math-classes-8.19.0/theory/naturals.v000066400000000000000000000203111460576051100175310ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.implementations.peano_naturals MathClasses.theory.rings Coq.Arith.PeanoNat MathClasses.categories.varieties MathClasses.theory.ua_transference. Require Export MathClasses.interfaces.naturals. Lemma to_semiring_involutive N `{Naturals N} N2 `{Naturals N2} x : naturals_to_semiring N2 N (naturals_to_semiring N N2 x) = x. Proof. rapply (proj2 (@categories.initials_unique' (varieties.Object semirings.theory) _ _ _ _ _ (semirings.object N) (semirings.object N2) _ naturals_initial _ naturals_initial) tt x). (* todo: separate pose necessary due to Coq bug *) Qed. Lemma to_semiring_unique `{Naturals N} `{SemiRing SR} (f: N → SR) `{!SemiRing_Morphism f} x : f x = naturals_to_semiring N SR x. Proof. symmetry. pose proof (@semirings.mor_from_sr_to_alg _ _ _ (semirings.variety N) _ _ _ (semirings.variety SR) (λ _, f) _). set (@varieties.arrow semirings.theory _ _ _ (semirings.variety N) _ _ _ (semirings.variety SR) (λ _, f) _). apply (naturals_initial _ a tt x). Qed. Lemma to_semiring_unique_alt `{Naturals N} `{SemiRing SR} (f g: N → SR) `{!SemiRing_Morphism f} `{!SemiRing_Morphism g} x : f x = g x. Proof. now rewrite (to_semiring_unique f), (to_semiring_unique g). Qed. Lemma morphisms_involutive `{Naturals N} `{SemiRing R} (f : R → N) (g : N → R) `{!SemiRing_Morphism f} `{!SemiRing_Morphism g} x : f (g x) = x. Proof. now apply (to_semiring_unique_alt (f ∘ g) id). Qed. Lemma to_semiring_twice `{Naturals N} `{SemiRing R1} `{SemiRing R2} (f : R1 → R2) (g : N → R1) (h : N → R2) `{!SemiRing_Morphism f} `{!SemiRing_Morphism g} `{!SemiRing_Morphism h} x : f (g x) = h x. Proof. now apply (to_semiring_unique_alt (f ∘ g) h). Qed. Lemma to_semiring_self `{Naturals N} (f : N → N) `{!SemiRing_Morphism f} x : f x = x. Proof. now apply (to_semiring_unique_alt f id). Qed. Lemma to_semiring_injective `{Naturals N} `{SemiRing A} (f: A → N) (g: N → A) `{!SemiRing_Morphism f} `{!SemiRing_Morphism g}: Injective g. Proof. repeat (split; try apply _). intros x y E. now rewrite <-(to_semiring_twice f g id x), <-(to_semiring_twice f g id y), E. Qed. #[global] Instance naturals_to_naturals_injective `{Naturals N} `{Naturals N2} (f: N → N2) `{!SemiRing_Morphism f}: Injective f | 15. Proof. now apply (to_semiring_injective (naturals_to_semiring N2 N) _). Qed. Section retract_is_nat. Context `{Naturals N} `{SemiRing SR}. Context (f : N → SR) `{inv_f : !Inverse f} `{!Surjective f} `{!SemiRing_Morphism f} `{!SemiRing_Morphism (f⁻¹)}. (* If we make this an instance, instance resolution will loop *) Definition retract_is_nat_to_sr : NaturalsToSemiRing SR := λ R _ _ _ _ , naturals_to_semiring N R ∘ f⁻¹. Section for_another_semirings. Context `{SemiRing R}. Instance: SemiRing_Morphism (naturals_to_semiring N R ∘ f⁻¹) := {}. Context (h : SR → R) `{!SemiRing_Morphism h}. Lemma same_morphism: naturals_to_semiring N R ∘ f⁻¹ = h. Proof. intros x y F. rewrite <-F. transitivity ((h ∘ (f ∘ f⁻¹)) x). symmetry. apply (to_semiring_unique (h ∘ f)). unfold compose. now rewrite jections.surjective_applied. Qed. End for_another_semirings. (* If we make this an instance, instance resolution will loop *) Program Instance retract_is_nat: Naturals SR (U:=retract_is_nat_to_sr). Next Obligation. unfold naturals_to_semiring, retract_is_nat_to_sr. apply _. Qed. Next Obligation. apply natural_initial. intros. now apply same_morphism. Qed. End retract_is_nat. Section contents. Context `{Naturals N}. Section borrowed_from_nat. Import universal_algebra. Import notations. Lemma induction (P: N → Prop) `{!Proper ((=) ==> iff) P}: P 0 → (∀ n, P n → P (1 + n)) → ∀ n, P n. Proof. intros. rewrite <-(to_semiring_involutive _ nat n). generalize (naturals_to_semiring N nat n). clear n. apply nat_induction. now rewrite preserves_0. intros n. rewrite preserves_plus, preserves_1. auto. Qed. Global Instance: Biinduction N. Proof. repeat intro. apply induction; firstorder. Qed. Lemma from_nat_stmt: ∀ (s: Statement varieties.semirings.theory) (w : Vars varieties.semirings.theory (varieties.semirings.object N) nat), (∀ v: Vars varieties.semirings.theory (varieties.semirings.object nat) nat, eval_stmt varieties.semirings.theory v s) → eval_stmt varieties.semirings.theory w s. Proof. pose proof (@naturals_initial nat _ _ _ _ _ _ _) as AI. pose proof (@naturals_initial N _ _ _ _ _ _ _) as BI. intros s w ?. apply (transfer_statement _ (@categories.initials_unique' semirings.Object _ _ _ _ _ (semirings.object nat) (semirings.object N) _ AI _ BI)). intuition. Qed. Let three_vars (x y z : N) (_: unit) v := match v with 0%nat => x | 1%nat => y | _ => z end. Let two_vars (x y : N) (_: unit) v := match v with 0%nat => x | _ => y end. Let no_vars (_: unit) (v: nat) := 0:N. Local Notation x' := (Var varieties.semirings.sig _ 0 tt). Local Notation y' := (Var varieties.semirings.sig _ 1 tt). Local Notation z' := (Var varieties.semirings.sig _ 2%nat tt). (* Some clever autoquoting tactic might make what follows even more automatic. *) (* The ugly [pose proof ... . apply that_thing.]'s are because of Coq bug 2185. *) Global Instance: ∀ z : N, LeftCancellation (+) z. Proof. intros x y z. rapply (from_nat_stmt (x' + y' === x' + z' -=> y' === z') (three_vars x y z)). intro. simpl. apply Nat.add_cancel_l. Qed. Global Instance: ∀ z : N, RightCancellation (+) z. Proof. intro. apply (right_cancel_from_left (+)). Qed. Global Instance: ∀ z : N, PropHolds (z ≠ 0) → LeftCancellation (.*.) z. Proof. intros z E x y. rapply (from_nat_stmt ((z' === 0 -=> Ext _ False) -=> z' * x' === z' * y' -=> x' === y') (three_vars x y z)). intro. simpl. intros. now apply (left_cancellation_ne_0 (.*.) (v () 2)). easy. Qed. Global Instance: ∀ z : N, PropHolds (z ≠ 0) → RightCancellation (.*.) z. Proof. intros ? ?. apply (right_cancel_from_left (.*.)). Qed. Instance nat_nontrivial: PropHolds ((1:N) ≠ 0). Proof. now rapply (from_nat_stmt (1 === 0 -=> Ext _ False) no_vars). Qed. Instance nat_nontrivial_apart `{Apart N} `{!TrivialApart N} : PropHolds ((1:N) ≶ 0). Proof. apply strong_setoids.ne_apart. solve_propholds. Qed. Lemma zero_sum (x y : N) : x + y = 0 → x = 0 ∧ y = 0. Proof. rapply (from_nat_stmt (x' + y' === 0 -=> Conj _ (x' === 0) (y' === 0)) (two_vars x y)). intro. simpl. apply Nat.eq_add_0. Qed. Lemma one_sum (x y : N) : x + y = 1 → (x = 1 ∧ y = 0) ∨ (x = 0 ∧ y = 1). Proof. rapply (from_nat_stmt (x' + y' === 1 -=> Disj _ (Conj _ (x' === 1) (y' === 0)) (Conj _ (x' === 0) (y' === 1))) (two_vars x y)). intros. simpl. intros. edestruct Nat.eq_add_1; eauto. Qed. Global Instance: ZeroProduct N. Proof. intros x y. rapply (from_nat_stmt (x' * y' === 0 -=>Disj _ (x' === 0) (y' === 0)) (two_vars x y)). intros ? E. destruct ((proj1 (Nat.eq_mul_0 _ _)) E); red; intuition. Qed. End borrowed_from_nat. Lemma nat_1_plus_ne_0 x : 1 + x ≠ 0. Proof. intro E. destruct (zero_sum 1 x E). now apply nat_nontrivial. Qed. Global Program Instance: ∀ x y: N, Decision (x = y) | 10 := λ x y, match decide (naturals_to_semiring _ nat x = naturals_to_semiring _ nat y) with | left E => left _ | right E => right _ end. Next Obligation. now rewrite <-(to_semiring_involutive _ nat x), <-(to_semiring_involutive _ nat y), E. Qed. Section with_a_ring. Context `{Ring R} `{!SemiRing_Morphism (f : N → R)} `{!Injective f}. Lemma to_ring_zero_sum x y : -f x = f y → x = 0 ∧ y = 0. Proof. intros E. apply zero_sum, (injective f). rewrite rings.preserves_0, rings.preserves_plus, <-E. now apply plus_negate_r. Qed. Lemma negate_to_ring x y : -f x = f y → f x = f y. Proof. intros E. destruct (to_ring_zero_sum x y E) as [E2 E3]. now rewrite E2, E3. Qed. End with_a_ring. End contents. (* Due to bug #2528 *) #[global] Hint Extern 6 (PropHolds (1 ≠ 0)) => eapply @nat_nontrivial : typeclass_instances. #[global] Hint Extern 6 (PropHolds (1 ≶ 0)) => eapply @nat_nontrivial_apart : typeclass_instances. math-classes-8.19.0/theory/products.v000066400000000000000000000035521460576051100175530ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra. Definition prod_equiv `{Equiv A} `{Equiv B} : Equiv (A * B) := λ p q, fst p = fst q ∧ snd p = snd q. (* Avoid eager application *) #[global] Hint Extern 0 (Equiv (_ * _)) => eapply @prod_equiv : typeclass_instances. Section product. Context `{Setoid A} `{Setoid B}. Global Instance: Setoid (A * B). Proof. firstorder auto. Qed. Global Instance pair_proper: Proper ((=) ==> (=) ==> (=)) (@pair A B). Proof. firstorder. Qed. Global Instance: Setoid_Morphism (@fst A B). Proof. constructor; try apply _. firstorder. Qed. Global Instance: Setoid_Morphism (@snd A B). Proof. constructor; try apply _. firstorder. Qed. Context `(A_dec : ∀ x y : A, Decision (x = y)) `(B_dec : ∀ x y : B, Decision (x = y)). Global Program Instance prod_dec: ∀ x y : A * B, Decision (x = y) := λ x y, match A_dec (fst x) (fst y) with | left _ => match B_dec (snd x) (snd y) with left _ => left _ | right _ => right _ end | right _ => right _ end. Solve Obligations with (program_simpl; firstorder). End product. Definition prod_fst_equiv X A `{Equiv X} : relation (X * A) := λ x y, fst x = fst y. Section product_fst. Context `{Setoid X}. Global Instance : `{Equivalence (prod_fst_equiv X A)}. Proof. unfold prod_fst_equiv. firstorder auto. Qed. End product_fst. Section dep_product. Context (I : Type) (c : I → Type) `{∀ i, Equiv (c i)} `{∀ i, Setoid (c i)}. Let dep_prod: Type := ∀ i, c i. Instance dep_prod_equiv: Equiv dep_prod := λ x y, ∀ i, x i = y i. Global Instance: Setoid dep_prod. Proof. constructor. repeat intro. reflexivity. intros ?? E ?. symmetry. apply E. intros ? y ??? i. transitivity (y i); firstorder. Qed. Global Instance dep_prod_morphism i : Setoid_Morphism (λ c: dep_prod, c i). Proof. firstorder auto. Qed. End dep_product. math-classes-8.19.0/theory/quote_monoid.v000066400000000000000000000104131460576051100204040ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra. Require MathClasses.interfaces.universal_algebra. Module ua := universal_algebra. Require MathClasses.varieties.monoids. Require MathClasses.theory.monoid_normalization. Notation msig := varieties.monoids.sig. Notation Op := (ua.Op msig False). Notation App := (ua.App msig False _ _). Notation Term V := (ua.Term0 msig V tt). Section contents. Context `{Monoid M}. Definition Vars V := V → M. Definition novars: Vars False := False_rect _. Definition singlevar (x: M): Vars unit := fun _ => x. Definition merge {A B} (a: Vars A) (b: Vars B): Vars (A+B) := fun i => match i with inl j => a j | inr j => b j end. Section Lookup. Class Lookup {A} (x: M) (f: Vars A) := { lookup: A; lookup_correct: f lookup ≡ x }. Global Arguments lookup {A} _ _ {Lookup}. Context (x: M) {A B} (va: Vars A) (vb: Vars B). Global Instance lookup_left `{!Lookup x va}: Lookup x (merge va vb). Proof. refine {| lookup := inl (lookup x va) |}. apply lookup_correct. Defined. Global Instance lookup_right `{!Lookup x vb}: Lookup x (merge va vb). Proof. refine {| lookup := inr (lookup x vb) |}. apply lookup_correct. Defined. Global Program Instance: Lookup x (singlevar x) := { lookup := tt }. End Lookup. Instance: `{MonUnit (Term V)} := λ V, ua.Op msig _ monoids.one. Instance: `{SgOp (Term V)} := λ V x, ua.App msig _ _ _ (ua.App msig _ _ _ (ua.Op msig _ monoids.mult) x). Notation eval V vs := (ua.eval _ (λ _, (vs: Vars V))). Section Quote. Class Quote {V} (l: Vars V) (n: M) {V'} (r: Vars V') := { quote: Term (V + V') ; eval_quote: @eval (V+V') (merge l r) quote ≡ n }. Arguments quote {V l} _ {V' r Quote}. Global Program Instance quote_zero V (v: Vars V): Quote v mon_unit novars := { quote := mon_unit }. Global Program Instance quote_mult V (v: Vars V) n V' (v': Vars V') m V'' (v'': Vars V'') `{!Quote v n v'} `{!Quote (merge v v') m v''}: Quote v (n & m) (merge v' v'') := { quote := ua.map_var msig (obvious: (V+V')→(V+(V'+V''))) (quote n) & ua.map_var msig (obvious:((V+V')+V'')→(V+(V'+V''))) (quote m) }. Next Obligation. Proof with auto. destruct Quote0, Quote1. subst. simpl. simpl. do 2 rewrite ua.eval_map_var. unfold ua_basic.algebra_op. simpl. f_equal; apply ua.eval_strong_proper; auto; repeat intro; intuition. Qed. Global Program Instance quote_old_var V (v: Vars V) x {i: Lookup x v}: Quote v x novars | 8 := { quote := ua.Var msig _ (inl (lookup x v)) tt }. Next Obligation. Proof. apply lookup_correct. Qed. Global Program Instance quote_new_var V (v: Vars V) x: Quote v x (singlevar x) | 9 := { quote := ua.Var msig _ (inr tt) tt }. End Quote. Definition quote': ∀ x {V'} {v: Vars V'} {d: Quote novars x v}, Term _ := @quote _ _. Definition eval_quote': ∀ x {V'} {v: Vars V'} {d: Quote novars x v}, eval _ (merge novars v) quote ≡ x := @eval_quote _ _. Arguments quote' _ {V' v d}. Arguments eval_quote' _ {V' v d}. Lemma quote_equality `{v: Vars V} `{v': Vars V'} (l r: M) `{!Quote novars l v} `{!Quote v r v'}: let heap := (merge v v') in eval _ heap (ua.map_var msig (obvious: False + V → V + V') quote) = eval _ heap quote → l = r. Proof with intuition. intros heap. destruct Quote0, Quote1. subst. subst heap. simpl. intro. rewrite <- H0. rewrite ua.eval_map_var. cut (eval _ (merge novars v) quote0 ≡ ua.eval msig (λ _, merge v v' ∘ obvious) quote0). intro E. simpl in *. rewrite E. reflexivity. apply (@ua.eval_strong_proper msig (λ _, M) _ _ (ne_list.one tt))... repeat intro... Qed. Lemma equal_by_normal `{v: Vars V} `{v': Vars V'} (l r: M) `{!Quote novars l v} `{!Quote v r v'}: ` (monoid_normalization.simplify _ (ua.map_var msig (obvious: False + V → V + V') quote)) ≡ ` (monoid_normalization.simplify _ quote) → l = r. Proof with intuition. do 2 destruct monoid_normalization.simplify. simpl. intros. subst. apply (quote_equality l r) . rewrite <- e, <- e0... Qed. Ltac monoid := apply (equal_by_normal _ _); vm_compute; reflexivity. Example ex x y z: x & (y & z) & mon_unit = mon_unit & (x & y) & z. Proof. monoid. Qed. End contents. math-classes-8.19.0/theory/rationals.v000066400000000000000000000204421460576051100177010ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.integers MathClasses.interfaces.naturals MathClasses.interfaces.rationals MathClasses.implementations.field_of_fractions MathClasses.implementations.natpair_integers MathClasses.theory.rings MathClasses.theory.integers MathClasses.theory.dec_fields. #[global] Program Instance slow_rat_dec `{Rationals Q} : ∀ x y: Q, Decision (x = y) | 10 := λ x y, match decide (rationals_to_frac Q (SRpair nat) x = rationals_to_frac Q (SRpair nat) y) with | left E => left _ | right E => right _ end. Next Obligation. now apply (injective (rationals_to_frac Q (SRpair nat))). Qed. Section another_integers. Context `{Rationals Q} `{Integers Z}. Add Ring Z : (stdlib_ring_theory Z). Notation ZtoQ := (integers_to_ring Z Q). Notation QtoFrac := (rationals_to_frac Q Z). Lemma rationals_decompose `{!SemiRing_Morphism (f : Z → Q)} (x : Q) : ∃ num, ∃ den, den ≠ 0 ∧ x = f num / f den. Proof. exists (num (QtoFrac x)), (den (QtoFrac x)). split. now apply den_ne_0. apply (injective QtoFrac). rewrite preserves_mult. rewrite preserves_dec_recip. change (QtoFrac x = (QtoFrac ∘ f) (num (QtoFrac x)) / (QtoFrac ∘ f) (den (QtoFrac x))). rewrite 2!(to_ring_unique_alt (QtoFrac ∘ f) Frac_inject). assert (Frac_inject (den (QtoFrac x)) ≠ 0) as P. apply injective_ne_0. apply den_ne_0. apply Frac_dec_mult_num_den. Qed. Global Instance integers_to_rationals_injective `{!SemiRing_Morphism (f: Z → Q)} : Injective f. Proof. split; try apply _. intros x y E. apply (injective (integers_to_ring Z Q)). now rewrite <-2!(to_ring_unique f). Qed. Context `{f : Q → Frac Z} `{!SemiRing_Morphism f}. Global Instance to_frac_inverse: Inverse f := λ x, ZtoQ (num x) / ZtoQ (den x). Global Instance: Surjective f. Proof. split; try apply _. intros x y E. unfold compose, inverse, to_frac_inverse, id. rewrite <-E. clear E. rewrite commutativity. apply (left_cancellation_ne_0 (.*.) (f (ZtoQ (den x)))). do 2 apply injective_ne_0. apply den_ne_0. rewrite <-preserves_mult, associativity. rewrite dec_recip_inverse, left_identity. change ((f ∘ ZtoQ) (num x) = (f ∘ ZtoQ) (den x) * x). rewrite 2!(to_ring_unique_alt (f ∘ ZtoQ) Frac_inject). unfold equiv, Frac_equiv. simpl. ring. apply injective_ne_0. now apply den_ne_0. Qed. (* Making this instance global results in loops *) Instance to_frac_bijective: Bijective f := {}. Global Instance to_frac_inverse_bijective: Bijective (f⁻¹) := _. Global Instance: SemiRing_Morphism (f⁻¹) := {}. End another_integers. Lemma to_frac_unique `{Rationals Q} `{Integers Z} (f g : Q → Frac Z) `{!SemiRing_Morphism f} `{!SemiRing_Morphism g} : ∀ x, f x = g x. Proof. intros x. apply (injective (g⁻¹)). change (f⁻¹ (f ⁻¹ ⁻¹ x) = g⁻¹ (g⁻¹ ⁻¹ x)). now rewrite 2!(jections.surjective_applied _). Qed. Definition rationals_to_rationals Q1 Q2 `{Rationals Q1} `{Rationals Q2} : Q1 → Q2 := (rationals_to_frac Q2 (SRpair nat))⁻¹ ∘ rationals_to_frac Q1 (SRpair nat). #[global] Hint Unfold rationals_to_rationals : typeclass_instances. Section another_rationals. Context `{Rationals Q1} `{Rationals Q2}. Hint Extern 4 => progress unfold rationals_to_rationals : typeclass_instances. Local Existing Instance to_frac_bijective. Global Instance: SemiRing_Morphism (rationals_to_rationals Q1 Q2) := _. Global Instance: Bijective (rationals_to_rationals Q1 Q2) := _. Global Instance: Bijective (rationals_to_rationals Q2 Q1) := _. Instance: Bijective (rationals_to_frac Q1 (SRpair nat)) := {}. Lemma to_rationals_involutive: ∀ x, rationals_to_rationals Q2 Q1 (rationals_to_rationals Q1 Q2 x) = x. Proof. intros x. unfold rationals_to_rationals, compose. rewrite (jections.surjective_applied _). apply (jections.bijective_applied _). Qed. Lemma to_rationals_unique (f : Q1 → Q2) `{!SemiRing_Morphism f} : ∀ x, f x = rationals_to_rationals Q1 Q2 x. Proof. intros x. apply (injective (rationals_to_rationals Q2 Q1)). rewrite to_rationals_involutive. change ((rationals_to_frac Q1 (SRpair nat)⁻¹) ((rationals_to_frac Q2 (SRpair nat) ∘ f) x) = x). rewrite (to_frac_unique (rationals_to_frac Q2 (SRpair nat) ∘ f) (rationals_to_frac Q1 (SRpair nat))). apply (jections.bijective_applied _). Qed. End another_rationals. Lemma to_rationals_unique_alt `{Rationals Q1} `{Rationals Q2} (f g : Q1 → Q2) `{!SemiRing_Morphism f} `{!SemiRing_Morphism g} : ∀ x, f x = g x. Proof. intros x. now rewrite (to_rationals_unique f), (to_rationals_unique g). Qed. Lemma morphisms_involutive `{Rationals Q1} `{Rationals Q2} (f : Q1 → Q2) (g : Q2 → Q1) `{!SemiRing_Morphism f} `{!SemiRing_Morphism g} : ∀ x, f (g x) = x. Proof. intros x. rewrite (to_rationals_unique f), (to_rationals_unique g). apply to_rationals_involutive. Qed. Global Instance injective_nats `{Rationals Q} `{Naturals N} `{!SemiRing_Morphism (f : N → Q)} : Injective f. Proof. split; try apply _. intros x y E. apply (injective (naturals_to_semiring N (SRpair nat))). apply rationals_embed_ints. now rewrite 2!(naturals.to_semiring_unique_alt f (integers_to_ring (SRpair nat) Q ∘ naturals_to_semiring N (SRpair nat))) in E. Qed. Section isomorphic_image_is_rationals. Context `{Rationals Q} `{DecField F}. Context (f : Q → F) `{!Inverse f} `{!Bijective f} `{!SemiRing_Morphism f}. Instance iso_to_frac: RationalsToFrac F := λ Z _ _ _ _ _ _ _ _, rationals_to_frac Q Z ∘ f⁻¹. Hint Extern 4 => progress unfold iso_to_frac : typeclass_instances. Instance: Bijective (f⁻¹) := _. Instance: SemiRing_Morphism (f⁻¹) := {}. Lemma iso_is_rationals: Rationals F. Proof. repeat (split; unfold rationals_to_frac; try apply _). intros x y E. apply (injective (f ∘ integers_to_ring Z Q)). now rewrite 2!(to_ring_unique (f ∘ integers_to_ring Z Q)). Qed. End isomorphic_image_is_rationals. Section alt_Build_Rationals. Context `{DecField F} `{Integers Z} (F_to_fracZ : F → Frac Z) `{!SemiRing_Morphism F_to_fracZ} `{!Injective F_to_fracZ}. Context (Z_to_F : Z → F) `{!SemiRing_Morphism Z_to_F} `{!Injective Z_to_F}. Program Instance alt_to_frac: RationalsToFrac F := λ B _ _ _ _ _ _ _ _ x, frac (integers_to_ring Z B (num (F_to_fracZ x))) (integers_to_ring Z B (den (F_to_fracZ x))) _. Next Obligation. apply injective_ne_0. apply den_ne_0. Qed. Instance: ∀ `{Integers B}, Proper ((=) ==> (=)) (rationals_to_frac F B). Proof. intros. intros ? ? E. unfold equiv, Frac_equiv. simpl. rewrite <-2!preserves_mult. apply sm_proper. change (F_to_fracZ x = F_to_fracZ y). now apply sm_proper. Qed. Instance: ∀ `{Integers B}, SemiRing_Morphism (rationals_to_frac F B). Proof. intros. repeat (split; try apply _); unfold equiv, Frac_equiv; simpl. intros x y. rewrite <-?preserves_mult, <-preserves_plus, <-preserves_mult. apply sm_proper. change (F_to_fracZ (x + y) = F_to_fracZ x + F_to_fracZ y). apply preserves_plus. rewrite <-(preserves_0 (f:=integers_to_ring Z B)), <-(preserves_1 (f:=integers_to_ring Z B)), <-2!preserves_mult. apply sm_proper. change (F_to_fracZ 0 = 0). apply preserves_0. intros x y. rewrite <-?preserves_mult. apply sm_proper. change (F_to_fracZ (x * y) = F_to_fracZ x * F_to_fracZ y). apply preserves_mult. rewrite <-(preserves_1 (f:=integers_to_ring Z B)), <-2!preserves_mult. apply sm_proper. change (F_to_fracZ 1 = 1). apply preserves_1. Qed. Instance: ∀ `{Integers B}, Injective (rationals_to_frac F B). Proof. intros. split; try apply _. intros x y E. unfold equiv, Frac_equiv in E. simpl in E. rewrite <-2!preserves_mult in E. apply (injective F_to_fracZ). now apply (injective (integers_to_ring Z B)). Qed. Instance: ∀ `{Integers B}, Injective (integers_to_ring B F). Proof. intros. split; try apply _. intros x y E. apply (injective (Z_to_F ∘ integers_to_ring B Z)). now rewrite 2!(to_ring_unique (Z_to_F ∘ integers_to_ring B Z)). Qed. Lemma alt_Build_Rationals: Rationals F. Proof. split; intros; apply _. Qed. End alt_Build_Rationals. math-classes-8.19.0/theory/ring_congruence.v000066400000000000000000000057461460576051100210660ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.theory.rings. Class RingCongruence A (R : relation A) `{Ring A} := { ring_congr_equivalence : Equivalence R ; ring_congr_subrelation : subrelation (=) R ; ring_congr_plus : Proper (R ==> R ==> R) (+) ; ring_congr_mult : Proper (R ==> R ==> R) (.*.) ; ring_congr_negate : Proper (R ==> R) (-)}. (* As far as I see, there are three ways to represent the quotient ring: - Define the congruence as a new Equiv instance. This leads to ambiguity. - Wrap it into a definition and define a ring structure on top of it. Definition Quotient A (R : relation A) := A. However, in order to avoid ambiguity we have to make it Opaque for type class resolution. - Wrap it into an inductive and define a ring structure on top of it. The latter makes the clearest distinction between the original structure and the quotient. Unfortunately, it is sometimes a bit verbose. *) Inductive Quotient A (R : relation A) := quotient_inject : Cast A (Quotient A R). #[global] Existing Instance quotient_inject. Arguments quotient_inject {A R} _. #[global] Instance quotient_rep {A R} : Cast (Quotient A R) A := λ x, match x with quotient_inject r => r end. Section quotient_ring. Context `{cong : RingCongruence A R}. Add Ring A : (rings.stdlib_ring_theory A). Existing Instance ring_congr_equivalence. Global Instance quotient_equiv: Equiv (Quotient A R) := λ x y, R ('x) ('y). Global Instance quotient_0: Zero (Quotient A R) := cast A (Quotient A R) 0. Global Instance quotient_1: One (Quotient A R) := cast A (Quotient A R) 1. Global Instance quotient_plus: Plus (Quotient A R) := λ x y, cast A (Quotient A R) ('x + 'y). Global Instance quotient_mult: Mult (Quotient A R) := λ x y, cast A (Quotient A R) ('x * 'y). Global Instance quotient_negate: Negate (Quotient A R) := λ x, cast A (Quotient A R) (-'x). Instance: Setoid (Quotient A R). Proof. constructor; unfold equiv, quotient_equiv. intros [x]. reflexivity. intros [x] [y] E. now symmetry. intros [x] [y] [z] E1 E2. now transitivity y. Qed. Instance: Proper ((=) ==> (=) ==> (=)) quotient_plus. Proof. intros [?] [?] E1 [?] [?] E2. now Tactics.rapply ring_congr_plus. Qed. Instance: Proper ((=) ==> (=) ==> (=)) quotient_mult. Proof. intros [?] [?] E1 [?] [?] E2. now Tactics.rapply ring_congr_mult. Qed. Instance: Proper ((=) ==> (=)) quotient_negate. Proof. intros [?] [?] E. now Tactics.rapply ring_congr_negate. Qed. Global Instance: Setoid_Morphism quotient_inject. Proof. split; try apply _. intros ?? E. now Tactics.rapply ring_congr_subrelation. Qed. Global Instance: Ring (Quotient A R). Proof. repeat (constructor; try apply _); repeat intros [?]; unfold mult, plus, sg_op, mon_unit, one_is_mon_unit, zero_is_mon_unit, plus_is_sg_op, mult_is_sg_op, quotient_0, negate, quotient_negate, quotient_1, quotient_plus, quotient_mult, cast in *; simpl; apply sm_proper; ring. Qed. End quotient_ring. math-classes-8.19.0/theory/ring_ideals.v000066400000000000000000000065711460576051100201740ustar00rootroot00000000000000 (* We define what a ring ideal is, show that they yield congruences, define what a kernel is, and show that kernels are ideal. *) Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.theory.rings. Require Export MathClasses.theory.ring_congruence. (* Require ua_congruence varieties.rings. *) Class RingIdeal A (P : A → Prop) `{Ring A} : Prop := { ideal_proper :: Proper ((=) ==> iff) P ; ideal_NonEmpty :: NonEmpty (sig P) ; ideal_closed_plus_negate : ∀ x y, P x → P y → P (x - y) ; ideal_closed_mult_r : ∀ x y, P x → P (x * y) ; ideal_closed_mult_l: ∀ x y, P y → P (x * y) }. Notation Factor A P := (Quotient A (λ x y, P (x - y))). Section ideal_congruence. Context `{ideal : RingIdeal A P}. Add Ring A2 : (rings.stdlib_ring_theory A). Let local_ideal_closed_plus_negate := (ideal_closed_plus_negate). Let local_ideal_closed_mult_l := (ideal_closed_mult_l). Let local_ideal_closed_mult_r := (ideal_closed_mult_r). (* If P is an ideal, we can easily derive some further closedness properties: *) Hint Resolve local_ideal_closed_plus_negate local_ideal_closed_mult_l local_ideal_closed_mult_r. Lemma ideal_closed_0 : P 0. Proof. destruct ideal_NonEmpty as [[x Px]]. rewrite <-(plus_negate_r x). intuition. Qed. Hint Resolve ideal_closed_0. Lemma ideal_closed_negate x : P x → P (-x). Proof. intros. rewrite <- rings.plus_0_l. intuition. Qed. Hint Resolve ideal_closed_negate. Lemma ideal_closed_plus x y : P x → P y → P (x + y). Proof. intros. assert (x + y = -(-x + -y)) as E by ring. rewrite E. intuition. Qed. Hint Resolve ideal_closed_plus. Global Instance: RingCongruence A (λ x y, P (x - y)). Proof. split. constructor. intros x. now rewrite plus_negate_r. intros x y E. rewrite negate_swap_r. intuition. intros x y z E1 E2. mc_setoid_replace (x - z) with ((x - y) + (y - z)) by ring. intuition. intros ?? E. now rewrite E, plus_negate_r. intros x1 x2 E1 y1 y2 E2. mc_setoid_replace (x1 + y1 - (x2 + y2)) with ((x1 - x2) + (y1 - y2)) by ring. intuition. intros x1 x2 E1 y1 y2 E2. mc_setoid_replace (x1 * y1 - (x2 * y2)) with ((x1 - x2) * y1 + x2 * (y1 - y2)) by ring. intuition. intros x1 x2 E. mc_setoid_replace (-x1 - - x2) with (-(x1 - x2)) by ring. intuition. Qed. Lemma factor_ring_eq (x y : Factor A P) : x = y ↔ P ('x - 'y). Proof. intuition. Qed. Lemma factor_ring_eq_0 (x y : Factor A P) : x = 0 ↔ P ('x). Proof. transitivity (P ('x - cast (Factor A P) A 0)). intuition. apply ideal_proper. unfold cast. simpl. ring. Qed. (* Let hint := rings.encode_operations R. Instance: Congruence rings.sig (λ _, congr_equiv). Proof. constructor; intros; apply _. Qed. *) End ideal_congruence. Section kernel_is_ideal. Context `{Ring A} `{Ring B} `{f : A → B} `{!SemiRing_Morphism f}. Add Ring A3 : (rings.stdlib_ring_theory A). Add Ring B3 : (rings.stdlib_ring_theory B). Definition kernel : A → Prop := (= 0) ∘ f. Global Instance: RingIdeal A kernel. Proof with ring. unfold kernel, compose, flip. split. intros ? ? E. now rewrite E. split. exists (0:A). apply preserves_0. intros ?? E E'. rewrite preserves_plus, preserves_negate, E, E'... intros ?? E. rewrite preserves_mult, E... intros ?? E. rewrite preserves_mult, E... Qed. End kernel_is_ideal. math-classes-8.19.0/theory/rings.v000066400000000000000000000305161460576051100170320ustar00rootroot00000000000000Require MathClasses.varieties.monoids MathClasses.theory.groups MathClasses.theory.strong_setoids. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra. Definition is_ne_0 `(x : R) `{Equiv R} `{Zero R} `{p : PropHolds (x ≠ 0)} : x ≠ 0 := p. Definition is_nonneg `(x : R) `{Equiv R} `{Le R} `{Zero R} `{p : PropHolds (0 ≤ x)} : 0 ≤ x := p. Definition is_pos `(x : R) `{Equiv R} `{Lt R} `{Zero R} `{p : PropHolds (0 < x)} : 0 < x := p. Lemma stdlib_semiring_theory R `{SemiRing R} : Ring_theory.semi_ring_theory 0 1 (+) (.*.) (=). Proof. constructor; intros. apply left_identity. apply commutativity. apply associativity. apply left_identity. apply left_absorb. apply commutativity. apply associativity. rewrite commutativity, distribute_l. now setoid_rewrite commutativity at 2 3. Qed. (* We cannot apply [left_cancellation (.*.) z] directly in case we have no [PropHolds (0 ≠ z)] instance in the context. *) Section cancellation. Context `{Ae : Equiv A} (op : A → A → A) `{!Zero A}. Lemma left_cancellation_ne_0 `{∀ z, PropHolds (z ≠ 0) → LeftCancellation op z} z : z ≠ 0 → LeftCancellation op z. Proof. auto. Qed. Lemma right_cancellation_ne_0 `{∀ z, PropHolds (z ≠ 0) → RightCancellation op z} z : z ≠ 0 → RightCancellation op z. Proof. auto. Qed. Lemma right_cancel_from_left `{!Setoid A} `{!Commutative op} `{!LeftCancellation op z} : RightCancellation op z. Proof. intros x y E. apply (left_cancellation op z). now rewrite 2!(commutativity z _). Qed. End cancellation. Section strong_cancellation. Context `{StrongSetoid A} (op : A → A → A). Lemma strong_right_cancel_from_left `{!Commutative op} `{!StrongLeftCancellation op z} : StrongRightCancellation op z. Proof. intros x y E. rewrite 2!(commutativity _ z). now apply (strong_left_cancellation op z). Qed. Global Instance strong_left_cancellation_cancel `{!StrongLeftCancellation op z} : LeftCancellation op z | 20. Proof. intros x y. rewrite <-!tight_apart. intros E1 E2. destruct E1. now apply (strong_left_cancellation op). Qed. Global Instance strong_right_cancellation_cancel `{!StrongRightCancellation op z} : RightCancellation op z | 20. Proof. intros x y. rewrite <-!tight_apart. intros E1 E2. destruct E1. now apply (strong_right_cancellation op). Qed. End strong_cancellation. Section semiring_props. Context `{SemiRing R}. Add Ring SR : (stdlib_semiring_theory R). Instance mult_ne_0 `{!NoZeroDivisors R} x y : PropHolds (x ≠ 0) → PropHolds (y ≠ 0) → PropHolds (x * y ≠ 0). Proof. intros Ex Ey Exy. unfold PropHolds in *. apply (no_zero_divisors x); split; eauto. Qed. Global Instance plus_0_r: RightIdentity (+) 0 := right_identity. Global Instance plus_0_l: LeftIdentity (+) 0 := left_identity. Global Instance mult_1_l: LeftIdentity (.*.) 1 := left_identity. Global Instance mult_1_r: RightIdentity (.*.) 1 := right_identity. Global Instance plus_assoc: Associative (+) := simple_associativity. Global Instance mult_assoc: Associative (.*.) := simple_associativity. Global Instance plus_comm: Commutative (+) := commutativity. Global Instance mult_comm: Commutative (.*.) := commutativity. Global Instance mult_0_l: LeftAbsorb (.*.) 0 := left_absorb. Global Instance mult_0_r: RightAbsorb (.*.) 0. Proof. intro. ring. Qed. Global Instance: RightDistribute (.*.) (+). Proof. intros x y z. ring. Qed. Lemma plus_mult_distr_r x y z: (x + y) * z = x * z + y * z. Proof. ring. Qed. Lemma plus_mult_distr_l x y z: x * (y + z) = x * y + x * z. Proof. ring. Qed. Global Instance: ∀ r : R, @Monoid_Morphism R R _ _ (0:R) (0:R) (+) (+) (r *.). Proof. repeat (constructor; try apply _). apply distribute_l. apply right_absorb. Qed. End semiring_props. (* Due to bug #2528 *) #[global] Hint Extern 3 (PropHolds (_ * _ ≠ 0)) => eapply @mult_ne_0 : typeclass_instances. Section semiringmor_props. Context `{SemiRing_Morphism A B f}. Lemma preserves_0: f 0 = 0. Proof (preserves_mon_unit (f:=f)). Lemma preserves_1: f 1 = 1. Proof (preserves_mon_unit (f:=f)). Lemma preserves_mult: ∀ x y, f (x * y) = f x * f y. Proof. intros. apply preserves_sg_op. Qed. Lemma preserves_plus: ∀ x y, f (x + y) = f x + f y. Proof. intros. apply preserves_sg_op. Qed. Instance: SemiRing B := semiringmor_b. Lemma preserves_2: f 2 = 2. Proof. rewrite preserves_plus. now rewrite preserves_1. Qed. Lemma preserves_3: f 3 = 3. Proof. now rewrite ?preserves_plus, ?preserves_1. Qed. Lemma preserves_4: f 4 = 4. Proof. now rewrite ?preserves_plus, ?preserves_1. Qed. Context `{!Injective f}. Instance injective_ne_0 x : PropHolds (x ≠ 0) → PropHolds (f x ≠ 0). Proof. intros. rewrite <-preserves_0. now apply (jections.injective_ne f). Qed. Lemma injective_ne_1 x : x ≠ 1 → f x ≠ 1. Proof. intros. rewrite <-preserves_1. now apply (jections.injective_ne f). Qed. End semiringmor_props. (* Due to bug #2528 *) #[global] Hint Extern 12 (PropHolds (_ _ ≠ 0)) => eapply @injective_ne_0 : typeclass_instances. Lemma stdlib_ring_theory R `{Ring R} : Ring_theory.ring_theory 0 1 (+) (.*.) (λ x y, x - y) (-) (=). Proof. constructor; intros. apply left_identity. apply commutativity. apply associativity. apply left_identity. apply commutativity. apply associativity. rewrite commutativity, distribute_l. now setoid_rewrite commutativity at 2 3. reflexivity. apply (right_inverse x). Qed. Section ring_props. Context `{Ring R}. Add Ring R: (stdlib_ring_theory R). Instance: LeftAbsorb (.*.) 0. Proof. intro. ring. Qed. Global Instance Ring_Semi: SemiRing R. Proof. repeat (constructor; try apply _). Qed. Definition negate_involutive x : - - x = x := groups.negate_involutive x. (* alias for convinience *) Lemma plus_negate_r x : x + -x = 0. Proof. ring. Qed. Lemma plus_negate_l x : -x + x = 0. Proof. ring. Qed. Lemma negate_swap_r x y : x - y = -(y - x). Proof. ring. Qed. Lemma negate_swap_l x y : -x + y = -(x - y). Proof. ring. Qed. Lemma negate_plus_distr x y : -(x + y) = -x + -y. Proof. ring. Qed. Lemma negate_mult x : -x = -1 * x. Proof. ring. Qed. Lemma negate_mult_distr_l x y : -(x * y) = -x * y. Proof. ring. Qed. Lemma negate_mult_distr_r x y : -(x * y) = x * -y. Proof. ring. Qed. Lemma negate_mult_negate x y : -x * -y = x * y. Proof. ring. Qed. Lemma negate_0: -0 = 0. Proof. ring. Qed. Global Instance minus_0_r: RightIdentity (λ x y, x - y) 0. Proof. intro x; rewrite negate_0; apply plus_0_r. Qed. Lemma equal_by_zero_sum x y : x - y = 0 ↔ x = y. Proof. split; intros E. rewrite <- (plus_0_l y). rewrite <- E. ring. rewrite E. ring. Qed. Lemma flip_negate x y : -x = y ↔ x = -y. Proof. split; intros E. now rewrite <-E, involutive. now rewrite E, involutive. Qed. Lemma flip_negate_0 x : -x = 0 ↔ x = 0. Proof. now rewrite flip_negate, negate_0. Qed. Lemma flip_negate_ne_0 x : -x ≠ 0 ↔ x ≠ 0. Proof. split; intros E ?; apply E; now apply flip_negate_0. Qed. Lemma negate_zero_prod_l x y : -x * y = 0 ↔ x * y = 0. Proof. split; intros E. apply (injective (-)). now rewrite negate_mult_distr_l, negate_0. apply (injective (-)). now rewrite negate_mult_distr_l, negate_involutive, negate_0. Qed. Lemma negate_zero_prod_r x y : x * -y = 0 ↔ x * y = 0. Proof. rewrite (commutativity x (-y)), (commutativity x y). apply negate_zero_prod_l. Qed. Lemma unit_no_zero_divisor (x : R) {unit : RingUnit x}: ¬ZeroDivisor x. Proof. destruct unit as [y x_y_unit]. intros [_ [z [z_nonzero xz_zero]]]. destruct z_nonzero. transitivity ((x * z) * y); try (rewrite xz_zero; ring). transitivity (x * y * z); try (rewrite x_y_unit; ring). ring. Qed. Context `{!NoZeroDivisors R} `{∀ x y, Stable (x = y)}. Global Instance mult_left_cancel: ∀ z, PropHolds (z ≠ 0) → LeftCancellation (.*.) z. Proof. intros z z_nonzero x y E. apply stable. intro U. apply (mult_ne_0 z (x - y) (is_ne_0 z)). intro. apply U. now apply equal_by_zero_sum. rewrite distribute_l, E. ring. Qed. Global Instance: ∀ z, PropHolds (z ≠ 0) → RightCancellation (.*.) z. Proof. intros ? ?. apply (right_cancel_from_left (.*.)). Qed. End ring_props. Section integral_domain_props. Context `{IntegralDomain R}. Instance intdom_nontrivial_apart `{Apart R} `{!TrivialApart R} : PropHolds (1 ≶ 0). Proof. apply strong_setoids.ne_apart. solve_propholds. Qed. End integral_domain_props. (* Due to bug #2528 *) #[global] Hint Extern 6 (PropHolds (1 ≶ 0)) => eapply @intdom_nontrivial_apart : typeclass_instances. Section ringmor_props. Context `{Ring A} `{Ring B} `{!SemiRing_Morphism (f : A → B)}. Definition preserves_negate x : f (-x) = -f x := groups.preserves_negate x. (* alias for convinience *) Lemma preserves_minus x y : f (x - y) = f x - f y. Proof. rewrite <-preserves_negate. apply preserves_plus. Qed. Lemma injective_preserves_0 : (∀ x, f x = 0 → x = 0) → Injective f. Proof. intros E1. split; try apply _. intros x y E. apply equal_by_zero_sum. apply E1. rewrite preserves_minus, E. now apply plus_negate_r. Qed. End ringmor_props. Section from_another_ring. Context `{Ring A} `{Setoid B} `{Bplus : Plus B} `{Zero B} `{Bmult : Mult B} `{One B} `{Bnegate : Negate B} (f : B → A) `{!Injective f} (plus_correct : ∀ x y, f (x + y) = f x + f y) (zero_correct : f 0 = 0) (mult_correct : ∀ x y, f (x * y) = f x * f y) (one_correct : f 1 = 1) (negate_correct : ∀ x, f (-x) = -f x). Lemma projected_ring: Ring B. Proof. split. now apply (groups.projected_ab_group f). now apply (groups.projected_com_monoid f mult_correct one_correct). repeat intro; apply (injective f). rewrite plus_correct, !mult_correct, plus_correct. now rewrite distribute_l. Qed. End from_another_ring. Section from_stdlib_semiring_theory. Context `(H: @semi_ring_theory R Rzero Rone Rplus Rmult Re) `{!@Setoid R Re} `{!Proper (Re ==> Re ==> Re) Rplus} `{!Proper (Re ==> Re ==> Re) Rmult}. Add Ring R2: H. Definition from_stdlib_semiring_theory: @SemiRing R Re Rplus Rmult Rzero Rone. Proof. repeat (constructor; try assumption); repeat intro ; unfold equiv, mon_unit, sg_op, zero_is_mon_unit, plus_is_sg_op, one_is_mon_unit, mult_is_sg_op, zero, mult, plus; ring. Qed. End from_stdlib_semiring_theory. Section from_stdlib_ring_theory. Context `(H: @ring_theory R Rzero Rone Rplus Rmult Rminus Rnegate Re) `{!@Setoid R Re} `{!Proper (Re ==> Re ==> Re) Rplus} `{!Proper (Re ==> Re ==> Re) Rmult} `{!Proper (Re ==> Re) Rnegate}. Add Ring R3: H. Definition from_stdlib_ring_theory: @Ring R Re Rplus Rmult Rzero Rone Rnegate. Proof. repeat (constructor; try assumption); repeat intro ; unfold equiv, mon_unit, sg_op, zero_is_mon_unit, plus_is_sg_op, one_is_mon_unit, mult_is_sg_op, mult, plus, negate; ring. Qed. End from_stdlib_ring_theory. #[global] Instance id_sr_morphism `{SemiRing A}: SemiRing_Morphism (@id A) := {}. Section morphism_composition. Context `{Equiv A} `{Mult A} `{Plus A} `{One A} `{Zero A} `{Equiv B} `{Mult B} `{Plus B} `{One B} `{Zero B} `{Equiv C} `{Mult C} `{Plus C} `{One C} `{Zero C} (f : A → B) (g : B → C). Instance compose_sr_morphism: SemiRing_Morphism f → SemiRing_Morphism g → SemiRing_Morphism (g ∘ f). Proof. split; try apply _; firstorder. Qed. Instance invert_sr_morphism: ∀ `{!Inverse f}, Bijective f → SemiRing_Morphism f → SemiRing_Morphism (f⁻¹). Proof. split; try apply _; firstorder. Qed. End morphism_composition. #[global] Hint Extern 4 (SemiRing_Morphism (_ ∘ _)) => class_apply @compose_sr_morphism : typeclass_instances. #[global] Hint Extern 4 (SemiRing_Morphism (_⁻¹)) => class_apply @invert_sr_morphism : typeclass_instances. #[global] Instance semiring_morphism_proper {A B eA eB pA mA zA oA pB mB zB oB} : Proper ((=) ==> (=)) (@SemiRing_Morphism A B eA eB pA mA zA oA pB mB zB oB). Proof. assert (∀ (f g : A → B), g = f → SemiRing_Morphism f → SemiRing_Morphism g) as P. intros f g E [? ? ? ?]. now split; try apply _; eapply groups.monoid_morphism_proper; eauto. intros f g ?; split; intros Mor. apply P with f. destruct Mor. now symmetry. now apply _. now apply P with g. Qed. math-classes-8.19.0/theory/sequences.v000066400000000000000000000064451460576051100177070ustar00rootroot00000000000000Require Import MathClasses.theory.categories MathClasses.interfaces.functors MathClasses.interfaces.abstract_algebra MathClasses.interfaces.sequences. #[global] Instance: Params (@extend) 6 := {}. Section contents. Context `{Sequence sq}. (* Some derived properties about inject, extend, and fmap: *) Lemma inject_epi `{Setoid A} `{Equiv B} `{SgOp B} `{MonUnit B} (f g: sq A → B) `{!Monoid_Morphism f} `{!Monoid_Morphism g} : f ∘ inject A = g ∘ inject A → f = g. Proof with intuition. intros. pose proof (setoidmor_a f). pose proof (monmor_b (f:=g)). pose proof (sequence_only_extend_commutes sq (f ∘ inject A) _) as E. pose proof (_ : Setoid_Morphism (f ∘ inject A)) as cm. rewrite (E f), (E g)... apply (sequence_extend_morphism sq)... apply cm. apply cm. Qed. Lemma extend_comp `{Equiv A} `{Equiv B} `{SgOp B} `{MonUnit B} `{Equiv C} `{SgOp C} `{MonUnit C} (f: B → C) (g: A → B) `{!Monoid_Morphism f} `{!Setoid_Morphism g} : extend (f ∘ g) (free:=sq) = f ∘ extend g (free:=sq). Proof with try apply _. intros. pose proof (setoidmor_a g). pose proof (monmor_a (f:=f)). pose proof (monmor_b (f:=f)). symmetry. apply (sequence_only_extend_commutes sq (f ∘ g))... symmetry. rewrite <- (sequence_extend_commutes sq g _) at 1. apply sm_proper. Qed. Lemma extend_inject `{Setoid A}: extend (inject A) = @id (sq A). Proof. symmetry. apply (sequence_only_extend_commutes sq); try apply _. apply sm_proper. Qed. (* Lemma fmap_alt `{Equiv A} `{Equiv B} (f: A → B) `{!Setoid_Morphism f} : extend (inject B ∘ f) = (fmap (v:=A) (w:=B) sq f: sq A → sq B). (* Remove (v:=A) (w:=B) *) Proof with try apply _. intros. pose proof (setoidmor_a f). pose proof (setoidmor_b f). symmetry. apply (sequence_only_extend_commutes sq (inject B ∘ f))... symmetry. apply (sequence_inject_natural sq)... Qed. *) Lemma fold_inject `{Monoid A}: fold sq ∘ inject A = id. Proof. apply (sequence_extend_commutes sq id). apply _. Qed. (* Lemma fold_map `{Setoid A} `{Monoid B} (f: A → B) `{!Setoid_Morphism f} : extend f (free:=sq) = fold sq ∘ fmap (v:=A) (w:=B) sq f. (* Remove (v:=A) (w:=B) *) Proof with try apply _. intros. symmetry. apply (sequence_only_extend_commutes sq _)... symmetry. change (f = extend id ∘ ((fmap sq f: sq A → sq B) ∘ inject A)). rewrite <- (sequence_inject_natural sq f _). change (f = fold sq ∘ inject B ∘ f). pose proof (_ : Morphisms.ProperProxy equiv f). rewrite fold_inject. rewrite compose_id_left. apply sm_proper. Qed. *) End contents. (* In the context of a SemiRing, we have two particularly useful folds: sum and product. *) Section semiring_folds. Context `{SemiRing R} `{Sequence sq}. Definition sum: sq R → R := @fold sq _ _ (0:R) (+). Definition product: sq R → R := @fold sq _ _ (1:R) mult. (* These are implicitly Monoid_Morphisms, and we also easily have: *) (* Lemma distribute_sum (a: R): (a *.) ∘ sum = sum ∘ fmap (v:=R) (w:=R) sq (a *.). Proof. unfold sum, fold. pose proof (_ : Monoid_Morphism (a *.)). rewrite <-(extend_comp (a *.) id). rewrite compose_id_right. rewrite (fold_map (a *.)). intros x y E. now rewrite E. Qed. *) End semiring_folds. math-classes-8.19.0/theory/series.v000066400000000000000000000240301460576051100171740ustar00rootroot00000000000000Require Import Coq.setoid_ring.Ring Coq.Arith.Factorial MathClasses.misc.workaround_tactics MathClasses.interfaces.abstract_algebra MathClasses.interfaces.additional_operations MathClasses.interfaces.orders MathClasses.interfaces.naturals MathClasses.interfaces.integers MathClasses.theory.nat_pow MathClasses.theory.int_pow MathClasses.theory.streams. Local Existing Instance srorder_semiring. Section series. Context `{SemiRingOrder A}. Add Ring A : (rings.stdlib_semiring_theory A). Add Ring nat : (rings.stdlib_semiring_theory nat). Class DecreasingNonNegative (s : ∞A) : Prop := decreasing_nonneg : ForAll (λ t, 0 ≤ hd (tl t) ≤ hd t) s. (* An equivalent definition is to say that given a point [s] in the stream, then every further point in that stream is in [[0,s]]. *) Lemma dnn_alt (s : ∞A) : DecreasingNonNegative s ↔ ForAll (λ s, ForAll (λ t, 0 ≤ hd t ≤ hd s) s) s. Proof. split; revert s; cofix FIX; intros s E. constructor. cut (hd s ≤ hd s); [|reflexivity]. set (x:=hd s). unfold x at 1. generalize s E x. clear s E x. cofix FIX2. intros s E. constructor. split; trivial. destruct E. now transitivity (hd (tl s)). destruct E. apply FIX2; trivial. now transitivity (hd s). destruct E. now apply FIX. constructor. now destruct E as [[? [? ?]] ?]. apply FIX. now destruct E. Qed. Definition dnn_Str_nth (s : ∞A) : DecreasingNonNegative s ↔ ∀ n, 0 ≤ Str_nth (S n) s ≤ Str_nth n s. Proof. split. intros dnn n. revert s dnn. induction n; intros s E; unfold Str_nth in *; simpl. apply E. apply IHn. now destruct E. revert s. cofix FIX. intros s E. constructor; unfold Str_nth in *; simpl in *. apply (E O). apply FIX. intros n. apply (E (S n)). Qed. Global Instance dnn_tl `(dnn : DecreasingNonNegative s) : DecreasingNonNegative (tl s). Proof. now destruct dnn. Qed. Global Instance dnn_Str_nth_tl `(dnn : DecreasingNonNegative s) : ∀ n, DecreasingNonNegative (Str_nth_tl n s). Proof. intros n. revert s dnn. induction n; [easy|]. intros s E. now apply IHn, _. Qed. Lemma dnn_hd_nonneg `(dnn : DecreasingNonNegative s) : 0 ≤ hd s. Proof. destruct dnn. now transitivity (hd (tl s)). Qed. Lemma dnn_Str_nth_nonneg `(dnn : DecreasingNonNegative s) n : 0 ≤ Str_nth n s. Proof. destruct n; apply (dnn_hd_nonneg _). Qed. Class IncreasingNonNegative (s : ∞A) : Prop := increasing_nonneg : ForAll (λ s, 0 ≤ hd s ≤ hd (tl s)) s. Lemma inn_Str_nth (s : ∞A) : IncreasingNonNegative s ↔ ∀ n, 0 ≤ Str_nth n s ≤ Str_nth (S n) s. Proof. split. intros E n. revert s E. induction n; intros s E; unfold Str_nth in *; simpl. apply E. apply IHn. now destruct E. revert s. cofix FIX. intros s E. constructor; unfold Str_nth in *; simpl in *. apply (E O). apply FIX. intros n. apply (E (S n)). Qed. Global Instance inn_tl `(inn : IncreasingNonNegative s) : IncreasingNonNegative (tl s). Proof. now destruct inn. Qed. Global Instance inn_Str_nth_tl `(inn : IncreasingNonNegative s) : ∀ n, IncreasingNonNegative (Str_nth_tl n s). Proof. intros n. revert s inn. induction n; trivial. intros s E. now apply IHn, _. Qed. Section every_other. CoFixpoint everyOther (s : ∞A) : ∞A := Cons (hd s) (everyOther (tl (tl s))). Lemma Str_nth_tl_everyOther (n : nat) (s : ∞A) : Str_nth_tl n (everyOther s) ≡ everyOther (Str_nth_tl (2 * n) s). Proof. revert s. induction n; intros s; simpl; trivial. rewrite IHn. replace (n + S (n + 0))%nat with (S (2*n))%nat. easy. change (1 + (1 + 1) * n = n + (1 + (n + 0))). ring. Qed. Lemma Str_nth_everyOther (n : nat) (s : ∞A) : Str_nth n (everyOther s) = Str_nth (2 * n) s. Proof. unfold Str_nth. rewrite Str_nth_tl_everyOther. reflexivity. Qed. Global Instance everyOther_dnn `(dnn : !DecreasingNonNegative s) : DecreasingNonNegative (everyOther s). Proof. revert s dnn. cofix FIX. intros s dnn. constructor; simpl. destruct dnn as [? [? ?]]. split. easy. now transitivity (hd (tl s)). now apply FIX, _. Qed. Global Instance everyOther_inc `(inn : !IncreasingNonNegative s) : IncreasingNonNegative (everyOther s). Proof. revert s inn. cofix FIX. intros s [? [? ?]]. constructor; simpl. split. easy. now transitivity (hd (tl s)). now apply FIX. Qed. End every_other. Section mult. Definition mult_Streams := zipWith (.*.). Global Instance mult_Streams_dnn `(dnn1 : !DecreasingNonNegative s1) `(dnn2 : !DecreasingNonNegative s2) : DecreasingNonNegative (mult_Streams s1 s2). Proof. revert s1 s2 dnn1 dnn2. cofix FIX. intros s1 s2 [[? ?] ?] [[? ?] ?]. constructor; simpl. split. now apply nonneg_mult_compat. now apply semirings.mult_le_compat. now apply FIX. Qed. Global Instance mult_Streams_inc `(inn1 : !IncreasingNonNegative s1) `(inn2 : !IncreasingNonNegative s2) : IncreasingNonNegative (mult_Streams s1 s2). Proof . revert s1 s2 inn1 inn2. cofix FIX. intros s1 s2 [? ?] [? ?]. constructor; simpl. split. now apply nonneg_mult_compat. now apply semirings.mult_le_compat. now apply FIX. Qed. Global Instance: Proper ((=) ==> (=) ==> (=)) mult_Streams. Proof. apply _. Qed. End mult. Section powers. Context (a : A). Definition powers_help : A → ∞A := iterate (.*a). Definition powers : ∞A := powers_help 1. Section with_nat_pow. Context `{Naturals N} `{!NatPowSpec A N pw} (f : nat → N) `{!SemiRing_Morphism f}. Lemma Str_nth_powers_help (n : nat) (c : A) : Str_nth n (powers_help c) = c * a ^ (f n). Proof. revert c. induction n using peano_naturals.nat_induction; intros c; unfold Str_nth in *; simpl. rewrite rings.preserves_0, nat_pow_0. ring. rewrite rings.preserves_plus, rings.preserves_1. rewrite nat_pow_S, IHn. ring. Qed. Lemma Str_nth_powers (n : nat) : Str_nth n powers = a ^ (f n). Proof. unfold powers. rewrite Str_nth_powers_help. ring. Qed. End with_nat_pow. Section with_int_pow. Context `{!Negate A} `{!DecRecip A} `{!DecField A} `{∀ x y : A, Decision (x = y)} `{Integers Z} `{!IntPowSpec A Z pw} (f : nat → Z) `{!SemiRing_Morphism f}. Lemma Str_nth_powers_help_int_pow (n : nat) (c : A) : Str_nth n (powers_help c) = c * a ^ (f n). Proof. rewrite (Str_nth_powers_help id). unfold id. now rewrite int_pow_nat_pow. Qed. Lemma Str_nth_powers_int_pow (n : nat) : Str_nth n powers = a ^ (f n). Proof. unfold powers. rewrite Str_nth_powers_help_int_pow. ring. Qed. End with_int_pow. End powers. Global Instance powers_help_proper: Proper ((=) ==> (=) ==> (=)) powers_help. Proof. intros ? ? E1 ? ? E2. apply stream_eq_Str_nth. intros n. now rewrite 2!(Str_nth_powers_help _ id), E1, E2. Qed. Global Instance powers_proper: Proper ((=) ==> (=)) powers. Proof. intros ? ? E. apply stream_eq_Str_nth. intros n. now rewrite 2!(Str_nth_powers _ id), E. Qed. Section increments. Context (d:A). Definition increments (x : A) : ∞A := iterate (+d) x. Lemma Str_nth_increments (n : nat) (x : A) : Str_nth n (increments x) = x + d * naturals_to_semiring nat A n. Proof. unfold increments. revert x. induction n using peano_naturals.nat_induction; intros c; unfold Str_nth in *; simpl. rewrite rings.preserves_0. now ring. rewrite IHn, rings.preserves_plus, rings.preserves_1. now ring. Qed. End increments. Section positives. Definition positives_help : A → ∞A := increments 1. Lemma Str_nth_positives_help (n : nat) (x : A) : Str_nth n (positives_help x) = x + naturals_to_semiring nat A n. Proof. unfold positives_help. rewrite Str_nth_increments. ring. Qed. Definition positives : ∞A := positives_help 1. Lemma Str_nth_positives (n : nat) : Str_nth n positives = 1 + naturals_to_semiring nat A n. Proof. unfold positives. now rewrite Str_nth_positives_help. Qed. End positives. Section factorials. CoFixpoint factorials_help (n c : A) : ∞A := c ::: factorials_help (1 + n) (n * c). Fixpoint fac_help (n : nat) (m c : A) : A := match n with | O => c | S n => (m + naturals_to_semiring nat A n) * fac_help n m c end. Lemma Str_nth_factorials_help (n : nat) (m c : A) : Str_nth n (factorials_help m c) = fac_help n m c. Proof. revert m c. induction n using peano_naturals.nat_induction; intros m c; unfold Str_nth in *; simpl. ring. rewrite IHn. clear IHn. induction n using peano_naturals.nat_induction; simpl. rewrite rings.preserves_0. now ring. rewrite IHn, rings.preserves_plus, rings.preserves_1. now ring. Qed. Definition factorials := factorials_help 1 1. Lemma Str_nth_factorials (n : nat) : Str_nth n factorials = naturals_to_semiring nat A (fact n). Proof. unfold factorials. rewrite Str_nth_factorials_help. induction n using peano_naturals.nat_induction; simpl. change (1 = naturals_to_semiring nat A 1). now rewrite rings.preserves_1. rewrite IHn. setoid_rewrite rings.preserves_plus. setoid_rewrite rings.preserves_mult. now ring. Qed. End factorials. End series. Section preservation. Context `{SemiRingOrder A} `{SemiRingOrder B} `{!SemiRing_Morphism (f : A → B)}. Lemma preserves_powers_help (a c : A) (n : nat) : f (Str_nth n (powers_help a c)) = Str_nth n (powers_help (f a) (f c)). Proof. rewrite 2!(Str_nth_powers_help _ id). rewrite rings.preserves_mult. now rewrite preserves_nat_pow. Qed. Lemma preserves_powers (a : A) (n : nat) : f (Str_nth n (powers a)) = Str_nth n (powers (f a)). Proof. rewrite 2!(Str_nth_powers _ id). apply preserves_nat_pow. Qed. Lemma preserves_positives (n : nat) : f (Str_nth n positives) = Str_nth n positives. Proof. rewrite 2!Str_nth_positives. rewrite rings.preserves_plus, rings.preserves_1. now rewrite <-(naturals.to_semiring_unique (f ∘ naturals_to_semiring nat A)). Qed. Lemma preserves_factorials (n : nat) : f (Str_nth n factorials) = Str_nth n factorials. Proof. rewrite 2!Str_nth_factorials. now rewrite <-(naturals.to_semiring_unique (f ∘ naturals_to_semiring nat A)). Qed. End preservation.math-classes-8.19.0/theory/setoids.v000066400000000000000000000052041460576051100173560ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.theory.products. Lemma ext_equiv_refl `{Setoid_Morphism A B f} : f = f. Proof. intros ?? E. pose proof (setoidmor_b f). now rewrite E. Qed. #[global] Instance ext_equiv_trans `{Equiv A} `{Equiv B} `{Reflexive (A:=A) (=)} `{Transitive (A:=B) (=)} : Transitive (_ : Equiv (A → B)). Proof. intros ? y ???? w ?. transitivity (y w); firstorder. Qed. #[global] Instance ext_equiv_sym `{Equiv A} `{Equiv B} `{Symmetric (A:=A) (=)} `{Symmetric (A:=B) (=)}: Symmetric (A:=A→B) (=). Proof. firstorder. Qed. Lemma ext_equiv_applied `{Setoid A} `{Equiv B} {f g : A → B} : f = g → ∀ x, f x = g x. Proof. intros E x. now apply E. Qed. Lemma ext_equiv_applied_iff `{Equiv A} `{Equiv B} `{!Setoid_Morphism (f : A → B)} `{!Setoid_Morphism (g : A → B)} : f = g ↔ ∀ x, f x = g x. Proof. pose proof (setoidmor_a f). pose proof (setoidmor_b f). split; intros E1. now apply ext_equiv_applied. intros x y E2. now rewrite E2. Qed. Lemma morphism_ne `{Equiv A} `{Equiv B} (f : A → B) `{!Setoid_Morphism f} x y : f x ≠ f y → x ≠ y. Proof. intros E1 E2. apply E1. now apply sm_proper. Qed. #[global] Instance: Equiv Prop := iff. #[global] Instance: Setoid Prop := {}. Lemma projected_setoid `{Setoid B} `{Equiv A} (f : A → B) (eq_correct : ∀ x y, x = y ↔ f x = f y) : Setoid A. Proof. constructor; repeat intro; apply eq_correct. reflexivity. symmetry; now apply eq_correct. transitivity (f y); now apply eq_correct. Qed. #[global] Instance sig_setoid `{Setoid A} (P : A → Prop) : Setoid (sig P). Proof. now apply (projected_setoid (@proj1_sig _ P)). Qed. #[global] Instance sigT_setoid `{Setoid A} (P : A → Type) : Setoid (sigT P). Proof. now apply (projected_setoid (@projT1 _ P)). Qed. Global Instance id_morphism `{Setoid A} : Setoid_Morphism (@id A). Proof. firstorder. Qed. Lemma compose_setoid_morphism `{Equiv A} `{Equiv B} `{Equiv C} (f : A → B) (g : B → C) : Setoid_Morphism f → Setoid_Morphism g → Setoid_Morphism (g ∘ f). Proof. firstorder. Qed. #[global] Hint Extern 4 (Setoid_Morphism (_ ∘ _)) => class_apply @compose_setoid_morphism : typeclass_instances. #[global] Instance morphism_proper `{Equiv A} `{Equiv B}: Proper ((=) ==> iff) (@Setoid_Morphism A B _ _). Proof. assert (∀ f g : A → B, f = g → Setoid_Morphism f → Setoid_Morphism g) as aux. intros f g E1 ?. pose proof (setoidmor_a f). pose proof (setoidmor_b f). split; try apply _. intros x y E2. now rewrite <-!(ext_equiv_applied E1 _), E2. intros f g; split; intros ?; eapply aux; eauto. pose proof (setoidmor_a g). pose proof (setoidmor_b g). now symmetry. Qed. math-classes-8.19.0/theory/shiftl.v000066400000000000000000000304211460576051100171740ustar00rootroot00000000000000Require MathClasses.orders.integers MathClasses.theory.dec_fields MathClasses.theory.nat_pow. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.naturals MathClasses.interfaces.integers MathClasses.interfaces.additional_operations MathClasses.interfaces.orders. Section shiftl. Context `{SemiRing A} `{!LeftCancellation (.*.) (2:A)} `{SemiRing B} `{!Biinduction B} `{!ShiftLSpec A B sl}. Add Ring A: (rings.stdlib_semiring_theory A). Add Ring B: (rings.stdlib_semiring_theory B). Global Instance: Proper ((=) ==> (=) ==> (=)) ((≪) : A → B → A) | 1. Proof shiftl_proper. Global Instance shiftl_mor_1: ∀ x : A, Setoid_Morphism (x≪) | 0. Proof. split; try apply _. Qed. Global Instance shiftl_mor_2: ∀ n : B, Setoid_Morphism (≪n) | 0. Proof. split; try apply _. solve_proper. Qed. Lemma shiftl_nat_pow_alt `{Naturals B2} `{!NatPowSpec A B2 pw} `{!SemiRing_Morphism (f : B2 → B)} x n : x ≪ f n = x * 2 ^ n. Proof. revert n. apply naturals.induction. solve_proper. rewrite rings.preserves_0, ?shiftl_0, nat_pow_0. ring. intros n E. rewrite rings.preserves_plus, rings.preserves_1, shiftl_S. rewrite E, nat_pow_S. ring. Qed. Lemma shiftl_nat_pow `{!NaturalsToSemiRing B} `{!Naturals B} `{!NatPowSpec A B np} x n : x ≪ n = x * 2 ^ n. Proof. change (x ≪ id n = x * 2 ^ n). apply shiftl_nat_pow_alt. Qed. Lemma shiftl_1 x : x ≪ (1:B) = 2 * x. Proof. now rewrite <-(rings.plus_0_r 1), shiftl_S, shiftl_0. Qed. Lemma shiftl_2 x : x ≪ (2:B) = 4 * x. Proof. rewrite shiftl_S, shiftl_1. ring. Qed. Global Instance shiftl_base_0: LeftAbsorb (≪) 0. Proof. intros n. pattern n. apply biinduction; clear n. solve_proper. now apply shiftl_0. intros n; split; intros E. rewrite shiftl_S, E. ring. apply (left_cancellation (.*.) 2). rewrite <-shiftl_S, E. ring. Qed. Lemma shiftl_exp_plus x n m : x ≪ (n + m) = x ≪ n ≪ m. Proof. pattern m. apply biinduction; clear m. solve_proper. now rewrite shiftl_0, rings.plus_0_r. intros m. setoid_replace (n + (1 + m)) with (1 + (n + m)) by ring. rewrite ?shiftl_S. split; intros E. now rewrite E. now apply (left_cancellation (.*.) 2). Qed. Lemma shiftl_order x n m: x ≪ n ≪ m = x ≪ m ≪ n. Proof. rewrite <-?shiftl_exp_plus. now rewrite commutativity. Qed. Lemma shiftl_reverse (x : A) (n m : B) : n + m = 0 → x ≪ n ≪ m = x. Proof. intros E. now rewrite <-shiftl_exp_plus, E, shiftl_0. Qed. Lemma shiftl_mult_l x y n : x * (y ≪ n) = (x * y) ≪ n. Proof. pattern n. apply biinduction; clear n. solve_proper. now rewrite ?shiftl_0. intros m. rewrite ?shiftl_S. split; intros E. rewrite <-E. ring. apply (left_cancellation (.*.) 2). rewrite <-E. ring. Qed. Lemma shiftl_mult_r x y n : (x ≪ n) * y = (x * y) ≪ n. Proof. now rewrite commutativity, shiftl_mult_l, commutativity. Qed. Lemma shiftl_base_plus x y n : (x + y) ≪ n = x ≪ n + y ≪ n. Proof. pattern n. apply biinduction; clear n. solve_proper. now rewrite ?shiftl_0. intros m. rewrite ?shiftl_S. split; intros E. rewrite E. ring. apply (left_cancellation (.*.) 2). rewrite E. ring. Qed. Lemma shiftl_base_nat_pow `{Naturals B2} `{!NatPowSpec A B2 pw} `{!SemiRing_Morphism (f : B2 → B)} x n m : (x ≪ n) ^ m = (x ^ m) ≪ (n * f m). Proof. revert m. apply naturals.induction. solve_proper. rewrite ?nat_pow_0. now rewrite rings.preserves_0, rings.mult_0_r, shiftl_0. intros m E. rewrite rings.preserves_plus, rings.preserves_1. rewrite rings.plus_mult_distr_l, rings.mult_1_r, shiftl_exp_plus. rewrite !nat_pow_S, E. now rewrite shiftl_mult_l, shiftl_mult_r. Qed. Lemma shiftl_negate `{Negate A} `{!Ring A} x n : (-x) ≪ n = -(x ≪ n). Proof. rewrite (rings.negate_mult x), (rings.negate_mult (x ≪ n)). symmetry. now apply shiftl_mult_l. Qed. Global Instance shiftl_inj: ∀ n, Injective (≪ n). Proof. repeat (split; try apply _). pattern n. apply biinduction; clear n. solve_proper. intros x y E. now rewrite ?shiftl_0 in E. intros n; split; intros E1 x y E2. apply E1. rewrite ?shiftl_S in E2. now apply (left_cancellation (.*.) 2). apply E1. now rewrite ?shiftl_S, E2. Qed. Instance shiftl_ne_0 x n : PropHolds (x ≠ 0) → PropHolds (x ≪ n ≠ 0). Proof. intros E1 E2. apply E1. apply (injective (≪ n)). now rewrite shiftl_base_0. Qed. Context `{Apart A} `{!FullPseudoSemiRingOrder (A:=A) Ale Alt} `{!PropHolds ((1:A) ≶ 0)}. Let shiftl_strict_order_embedding (x y : A) (n : B) : x < y ↔ x ≪ n < y ≪ n. Proof. revert n. apply (biinduction_iff (x < y) (λ n, x ≪ n < y ≪ n)). solve_proper. now rewrite 2!shiftl_0. intros n. rewrite !shiftl_S. split; intros E. now apply (strictly_order_preserving (2 *.)). now apply (strictly_order_reflecting (2 *.)). Qed. Global Instance: ∀ n, StrictOrderEmbedding (≪ n). Proof. repeat (split; try apply _); intros. now apply shiftl_strict_order_embedding. eapply shiftl_strict_order_embedding. eassumption. Qed. Global Instance: ∀ n, OrderEmbedding (≪ n). Proof. split. now apply maps.full_pseudo_order_preserving. now apply maps.full_pseudo_order_reflecting. Qed. Global Instance shiftl_strong_inj: ∀ n, StrongInjective (≪ n). Proof. intros. apply maps.pseudo_order_embedding_inj. Qed. Lemma shiftl_le_flip_r `{Negate B} `{!Ring B} (x y : A) (n : B) : x ≤ y ≪ (-n) ↔ x ≪ n ≤ y. Proof. split; intros E. apply (order_reflecting (≪ -n)). now rewrite shiftl_reverse by now apply rings.plus_negate_r. apply (order_reflecting (≪ n)). now rewrite shiftl_reverse by now apply rings.plus_negate_l. Qed. Lemma shiftl_le_flip_l `{Negate B} `{!Ring B} (x y : A) (n : B) : x ≪ (-n) ≤ y ↔ x ≤ y ≪ n. Proof. now rewrite <-shiftl_le_flip_r, rings.negate_involutive. Qed. Instance shiftl_nonneg (x : A) (n : B) : PropHolds (0 ≤ x) → PropHolds (0 ≤ x ≪ n). Proof. intro. rewrite <-(shiftl_base_0 n). now apply (order_preserving (≪ n)). Qed. Instance shiftl_pos (x : A) (n : B) : PropHolds (0 < x) → PropHolds (0 < x ≪ n). Proof. intro. rewrite <-(shiftl_base_0 n). now apply (strictly_order_preserving (≪ n)). Qed. End shiftl. (* Due to bug #2528 *) #[global] Hint Extern 18 (PropHolds (_ ≪ _ ≠ 0)) => eapply @shiftl_ne_0 : typeclass_instances. #[global] Hint Extern 18 (PropHolds (0 ≤ _ ≪ _)) => eapply @shiftl_nonneg : typeclass_instances. #[global] Hint Extern 18 (PropHolds (0 < _ ≪ _)) => eapply @shiftl_pos : typeclass_instances. Section preservation. Context `{SemiRing B} `{!Biinduction B} `{SemiRing A1} `{!ShiftLSpec A1 B sl1} `{SemiRing A2} `{!LeftCancellation (.*.) (2:A2)} `{!ShiftLSpec A2 B sl2} `{!SemiRing_Morphism (f : A1 → A2)}. Lemma preserves_shiftl x (n : B) : f (x ≪ n) = (f x) ≪ n. Proof. revert n. apply biinduction. solve_proper. now rewrite 2!shiftl_0. intros n; split; intros IH. rewrite 2!shiftl_S. now rewrite rings.preserves_mult, rings.preserves_2, IH. apply (left_cancellation (.*.) 2). rewrite <-(rings.preserves_2 (f:=f)) at 1. rewrite <-rings.preserves_mult, <-shiftl_S, IH. now rewrite shiftl_S. Qed. End preservation. Section exp_preservation. Context `{SemiRing B1} `{!Biinduction B1} `{SemiRing B2} `{!Biinduction B2} `{SemiRing A} `{!LeftCancellation (.*.) (2:A)} `{!ShiftLSpec A B1 sl1} `{!ShiftLSpec A B2 sl2} `{!SemiRing_Morphism (f : B1 → B2)}. Lemma preserves_shiftl_exp x (n : B1) : x ≪ f n = x ≪ n. Proof. revert n. apply biinduction. solve_proper. now rewrite rings.preserves_0, ?shiftl_0. intros n. rewrite rings.preserves_plus, rings.preserves_1, ?shiftl_S. split; intros E. now rewrite E. now apply (left_cancellation (.*.) 2). Qed. End exp_preservation. Section shiftl_dec_field. Context `{SemiRing R} `{Integers Z} `{!ShiftLSpec R Z sl} `{DecField F} `{∀ x y : F, Decision (x = y)} `{!PropHolds ((2:F) ≠ 0)} `{!IntPowSpec F Z ipw} `{!SemiRing_Morphism (f : R → F)}. Add Ring F: (rings.stdlib_ring_theory F). Add Ring Z: (rings.stdlib_ring_theory Z). Existing Instance int_pow_proper. Lemma shiftl_to_int_pow x n : f (x ≪ n) = f x * 2 ^ n. Proof. revert n. apply biinduction. solve_proper. now rewrite shiftl_0, int_pow_0, rings.mult_1_r. intros n. rewrite shiftl_S, int_pow_S by solve_propholds. rewrite rings.preserves_mult, rings.preserves_2. rewrite associativity, (commutativity (f x) 2), <-associativity. split; intros E. now rewrite E. now apply (left_cancellation (.*.) 2). Qed. Lemma shiftl_base_1_to_int_pow n : f (1 ≪ n) = 2 ^ n. Proof. now rewrite shiftl_to_int_pow, rings.preserves_1, rings.mult_1_l. Qed. Lemma shiftl_negate_1_to_half x : f (x ≪ -1) = f x / 2. Proof. rewrite shiftl_to_int_pow. apply (left_cancellation (.*.) 2). transitivity (f x * (2 * 2 ^ (-1))); [ring |]. transitivity (f x * (2 / 2)); [| ring]. rewrite dec_recip_inverse, <-int_pow_S by assumption. now rewrite rings.plus_negate_r, int_pow_0. Qed. Lemma shiftl_negate_1_to_fourth x : f (x ≪ -2) = f x / 4. Proof. rewrite shiftl_to_int_pow. apply (left_cancellation (.*.) (2 * 2)). transitivity (f x * (2 * (2 * 2 ^ (-2)))); [ring |]. transitivity (f x * (4 / 4)); [| ring]. assert ((4:F) ≠ 0). setoid_replace 4 with (2*2) by ring. solve_propholds. rewrite dec_recip_inverse, <-!int_pow_S by assumption. setoid_replace (1 + (1 - 2) : Z) with (0 : Z) by ring. now rewrite int_pow_0. Qed. End shiftl_dec_field. Section more_shiftl_dec_field. Context `{DecField A} `{Integers B} `{∀ x y : A, Decision (x = y)} `{!PropHolds ((2:A) ≠ 0)} `{!ShiftLSpec A B sl} `{!IntPowSpec A B ipw}. Lemma shiftl_int_pow x n : x ≪ n = x * 2 ^ n. Proof. change (id (x ≪ n) = id x * 2 ^ n). apply shiftl_to_int_pow. Qed. Lemma shiftl_base_1_int_pow n : 1 ≪ n = 2 ^ n. Proof. now rewrite shiftl_int_pow, rings.mult_1_l. Qed. Lemma shiftl_negate_1_half x : x ≪ (-1) = x / 2. Proof. change (id (x ≪ (-1)) = id x / 2). now apply shiftl_negate_1_to_half. Qed. Lemma shiftl_negate_1_fourth x : x ≪ (-2) = x / 4. Proof. change (id (x ≪ (-2)) = id x / 4). now apply shiftl_negate_1_to_fourth. Qed. End more_shiftl_dec_field. Section shiftl_field. Context `{Ring R} `{Integers Z} `{!ShiftLSpec R Z sl} `{Field F} `{!PropHolds ((2:F) ≶ 0)} `{Naturals N} `{!NatPowSpec F N npw} `{!SemiRing_Morphism (g : N → Z)} `{!SemiRing_Morphism (f : R → F)}. Add Ring F2: (rings.stdlib_ring_theory F). Add Ring Z2: (rings.stdlib_ring_theory Z). Lemma shiftl_negate_nat_pow x n : f (x ≪ (-g n)) * 2 ^ n = f x. Proof. pose proof nat_pow_proper. pattern n. apply naturals.induction; clear n. solve_proper. rewrite rings.preserves_0, rings.negate_0, shiftl_0. rewrite nat_pow_0. ring. intros n E. rewrite rings.preserves_plus, rings.preserves_1. etransitivity; [| eassumption]. setoid_replace (-g n) with (1 - (1 + g n)) by ring. rewrite shiftl_S, rings.preserves_mult, rings.preserves_2. rewrite nat_pow_S. ring. Qed. Lemma shiftl_negate_to_recip_nat_pow x n P2n : f (x ≪ (-g n)) = f x // (2 ^ n)↾P2n. Proof. apply (right_cancellation (.*.) (2 ^ n)). rewrite shiftl_negate_nat_pow. transitivity (f x * (2 ^ n // (2 ^ n)↾P2n)); [| ring]. rewrite fields.reciperse_alt. ring. Qed. End shiftl_field. Section default_shiftl_naturals. Context `{SemiRing A} `{Naturals B} `{!NatPowSpec A B pw}. Global Instance default_shiftl: ShiftL A B | 10 := λ x n, x * 2 ^ n. Global Instance: ShiftLSpec A B default_shiftl. Proof. now apply shiftl_spec_from_nat_pow. Qed. End default_shiftl_naturals. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Opaque default_shiftl. Set Warnings "+unsupported-attributes". Section default_shiftl_integers. Context `{DecField A} `{!PropHolds ((2:A) ≠ 0)} `{Integers B} `{!IntPowSpec A B ipw}. Global Instance default_shiftl_int: ShiftL A B | 9 := λ x n, x * 2 ^ n. Global Instance: ShiftLSpec A B default_shiftl_int. Proof. now apply shiftl_spec_from_int_pow. Qed. End default_shiftl_integers. Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Opaque default_shiftl_int. math-classes-8.19.0/theory/streams.v000066400000000000000000000112161460576051100173620ustar00rootroot00000000000000(* In the standard library equality on streams is defined as pointwise Leibniz equality. Here we prove similar results, but we use setoid equality instead. *) Require Export MathClasses.theory.CoqStreams. Require Import MathClasses.implementations.peano_naturals MathClasses.interfaces.abstract_algebra. Section streams. Context `{Setoid A}. CoInductive Stream_eq_coind (s1 s2: ∞A) : Prop := stream_eq_coind : hd s1 = hd s2 → Stream_eq_coind (tl s1) (tl s2) → Stream_eq_coind s1 s2. Global Instance stream_eq: Equiv (∞A) := Stream_eq_coind. Global Instance: Setoid (∞A). Proof. split. cofix FIX. now constructor. cofix FIX. intros ? ? E. constructor. symmetry. now destruct E. apply FIX. now destruct E. cofix FIX. intros ? ? ? E1 E2. constructor. transitivity (hd y). now destruct E1. now destruct E2. apply FIX with (tl y). now destruct E1. now destruct E2. Qed. Global Instance: Proper ((=) ==> (=)) (@Cons A). Proof. intros ? ? E. constructor. simpl. now rewrite E. easy. Qed. Global Instance: Proper ((=) ==> (=)) (@hd A). Proof. now intros ? ? []. Qed. Global Instance: Proper ((=) ==> (=)) (@tl A). Proof. now intros ? ? []. Qed. Lemma Str_nth_tl_S (s : ∞A) n : Str_nth_tl (S n) s ≡ tl (Str_nth_tl n s). Proof. now rewrite tl_nth_tl. Qed. Global Instance: Proper ((=) ==> (=) ==> (=)) (@Str_nth_tl A). Proof. intros n m E1 ? ? E2. unfold equiv, peano_naturals.nat_equiv in E1. rewrite E1. clear E1 n. induction m. easy. simpl. rewrite <-2!tl_nth_tl. now rewrite IHm. Qed. Global Instance: Proper ((=) ==> (=) ==> (=)) (@Str_nth A). Proof. intros ? ? E1 ? ? E2. unfold Str_nth. now rewrite E1, E2. Qed. Lemma stream_eq_Str_nth s1 s2 : s1 = s2 ↔ ∀ n, Str_nth n s1 = Str_nth n s2. Proof. split. intros E ?. now rewrite E. revert s1 s2. cofix FIX. intros s1 s2 E. constructor. now apply (E O). apply FIX. intros. now apply (E (S n)). Qed. Global Instance ForAll_proper `{!Proper ((=) ==> iff) (P : ∞A → Prop)} : Proper ((=) ==> iff) (ForAll P). Proof. assert (∀ x y, x = y → ForAll P x → ForAll P y). cofix FIX. intros ? ? E1 E2. constructor. rewrite <-E1. now destruct E2. apply FIX with (tl x). now rewrite E1. now destruct E2. intros ? ? E1. intuition eauto. symmetry in E1; intuition eauto. Qed. Lemma ForAll_tl (P : ∞A → Prop) s : ForAll P s → ForAll P (tl s). Proof. apply (ForAll_Str_nth_tl 1). Qed. Lemma ForAll_True (s : ∞A) : ForAll (λ _, True) s. Proof. revert s. cofix F. intros. constructor; trivial. Qed. Definition EventuallyForAll (P : ∞A → Prop) := ForAll (λ s, P s → P (tl s)). Lemma EventuallyForAll_tl P s : EventuallyForAll P s → EventuallyForAll P (tl s). Proof. repeat intro. now apply ForAll_tl. Qed. Lemma EventuallyForAll_Str_nth_tl P n s : EventuallyForAll P s → EventuallyForAll P (Str_nth_tl n s). Proof. revert s. induction n. easy. intros. now apply IHn, EventuallyForAll_tl. Qed. Global Instance EventuallyForAll_proper `{!Proper ((=) ==> iff) (P : ∞A → Prop)} : Proper ((=) ==> iff) (EventuallyForAll P). Proof. assert (Proper ((=) ==> iff) (λ s, P s → P (tl s))). intros ? ? E. now rewrite E. intros ? ? E. now rewrite E. Qed. CoFixpoint iterate (f:A → A) (x:A) : ∞A := x ::: iterate f (f x). CoFixpoint repeat (x:A) : ∞A := x ::: repeat x. End streams. Section more. Context `{Setoid A} `{Setoid B}. CoInductive ForAllIf (PA : ∞A → Prop) (PB : ∞B → Prop) : ∞A → ∞B → Prop := ext_if : ∀ s1 s2, (PA s1 → PB s2) → ForAllIf PA PB (tl s1) (tl s2) → ForAllIf PA PB s1 s2. Global Instance ForAllIf_proper `{!Proper ((=) ==> iff) (PA : ∞A → Prop)} `{!Proper ((=) ==> iff) (PB : ∞B → Prop)} : Proper ((=) ==> (=) ==> iff) (ForAllIf PA PB). Proof. assert (∀ x1 y1 x2 y2, x1 = y1 → x2 = y2 → ForAllIf PA PB x1 x2 → ForAllIf PA PB y1 y2) as P. cofix FIX. intros ? ? ? ? E1 E2 E3. constructor. rewrite <-E1, <-E2. now destruct E3. apply FIX with (tl x1) (tl x2). now rewrite E1. now rewrite E2. now destruct E3. repeat intro. now split; apply P. Qed. Global Instance map_proper `{!Proper ((=) ==> (=)) (f : A → B)} : Proper ((=) ==> (=)) (map f). Proof. cofix FIX. intros ? ? E. constructor. simpl. destruct E as [E]. now rewrite E. simpl. apply FIX. now apply E. Qed. Context `{Setoid C}. Global Instance zipWith_proper `{!Proper ((=) ==> (=) ==> (=)) (f : A → B → C)} : Proper ((=) ==> (=) ==> (=)) (zipWith f). Proof. cofix FIX. intros ? ? E1 ? ? E2. constructor. simpl. destruct E1 as [E1], E2 as [E2]. now rewrite E1, E2. simpl. apply FIX. now apply E1. now apply E2. Qed. End more. math-classes-8.19.0/theory/strong_setoids.v000066400000000000000000000174311460576051100207570ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.theory.jections. Section contents. Context `{StrongSetoid A}. Global Instance: Setoid A. Proof. split. intros x. rewrite <-tight_apart. now apply (irreflexivity (≶)). intros x y. rewrite <-?tight_apart. now apply not_symmetry. intros x y z. rewrite <-?tight_apart. intros E1 E2 E3. destruct (cotransitive E3 y); contradiction. Qed. Global Instance apart_proper: Proper ((=) ==> (=) ==> iff) (≶). Proof. assert (∀ x₁ y x₂, x₁ ≶ y → x₁ = x₂ → x₂ ≶ y) as P1. intros ? ? ? E Ex. destruct (cotransitive E x₂); trivial. apply tight_apart in Ex. destruct Ex. now symmetry. assert (∀ x₁ y₁ x₂ y₂, x₁ ≶ y₁ → x₁ = x₂ → y₁ = y₂ → x₂ ≶ y₂) as P2. intros ? ? ? ? E Ex Ey. apply P1 with x₁; trivial. symmetry. apply P1 with y₁; trivial. now symmetry. intros ? ? E1 ? ? E2. split; intro; eapply P2; eauto; now symmetry. Qed. Instance apart_ne x y : PropHolds (x ≶ y) → PropHolds (x ≠ y). Proof. firstorder. Qed. Global Instance: ∀ x y, Stable (x = y). Proof. intros x y. unfold Stable, DN. rewrite <-tight_apart. tauto. Qed. End contents. (* Due to bug #2528 *) #[global] Hint Extern 3 (PropHolds (_ ≠ _)) => eapply @apart_ne : typeclass_instances. Lemma projected_strong_setoid `{StrongSetoid B} `{Equiv A} `{Apart A} (f: A → B) (eq_correct : ∀ x y, x = y ↔ f x = f y) (apart_correct : ∀ x y, x ≶ y ↔ f x ≶ f y) : StrongSetoid A. Proof. split. intros x. red. rewrite apart_correct. apply (irreflexivity (≶)). intros x y. rewrite !apart_correct. now symmetry. intros x y E z. rewrite !apart_correct. apply cotransitive. now apply apart_correct. intros x y. rewrite apart_correct, eq_correct. now apply tight_apart. Qed. #[global] Instance sig_strong_setoid `{StrongSetoid A} (P: A → Prop): StrongSetoid (sig P). Proof. now apply (projected_strong_setoid (@proj1_sig _ P)). Qed. Section morphisms. Context `{Equiv A} `{Apart A} `{Equiv B} `{Apart B} `{Equiv C} `{Apart C}. Existing Instance strong_setoidmor_a. Existing Instance strong_setoidmor_b. Global Instance strong_morphism_proper `{!StrongSetoid_Morphism (f : A → B)} : Setoid_Morphism f | 10. Proof. split; try apply _. intros ? ?. rewrite <-!tight_apart. intros E1 E2. destruct E1. now apply (strong_extensionality f). Qed. Global Instance strong_injective_injective `{!StrongInjective (f : A → B)} : Injective f. Proof. pose proof (strong_injective_mor f). split; try apply _. intros ? ?. rewrite <-!tight_apart. intros E1 E2. destruct E1. now apply (strong_injective f). Qed. (* If a morphism satisfies the binary strong extensionality property, it is strongly extensional in both coordinates. *) Global Instance strong_setoid_morphism_1 `{!StrongSetoid_BinaryMorphism (f : A → B → C)} : ∀ z, StrongSetoid_Morphism (f z). Proof. pose proof (strong_binary_setoidmor_a f). pose proof (strong_binary_setoidmor_b f). pose proof (strong_binary_setoidmor_c f). intros z. split; try apply _. intros x y E. destruct (strong_binary_extensionality f z x z y); trivial. now destruct (irreflexivity (≶) z). Qed. Global Instance strong_setoid_morphism_unary_2 `{!StrongSetoid_BinaryMorphism (f : A → B → C)} : ∀ z, StrongSetoid_Morphism (λ x, f x z). Proof. pose proof (strong_binary_setoidmor_a f). pose proof (strong_binary_setoidmor_b f). pose proof (strong_binary_setoidmor_c f). intros z. split; try apply _. intros x y E. destruct (strong_binary_extensionality f x z y z); trivial. now destruct (irreflexivity (≶) z). Qed. (* Conversely, if a morphism is strongly extensional in both coordinates, it satisfies the binary strong extensionality property. We don't make this an instance in order to avoid loops. *) Lemma strong_binary_setoid_morphism_both_coordinates `{!StrongSetoid A} `{!StrongSetoid B} `{!StrongSetoid C} {f : A → B → C} `{∀ z, StrongSetoid_Morphism (f z)} `{∀ z, StrongSetoid_Morphism (λ x, f x z)} : StrongSetoid_BinaryMorphism f. Proof. split; try apply _. intros x₁ y₁ x₂ y₂ E. destruct (cotransitive E (f x₂ y₁)). left. now apply (strong_extensionality (λ x, f x y₁)). right. now apply (strong_extensionality (f x₂)). Qed. Global Instance binary_strong_morphism_proper `{!StrongSetoid_BinaryMorphism (f : A → B → C)} : Proper ((=) ==> (=) ==> (=)) f. Proof. pose proof (strong_binary_setoidmor_a f). pose proof (strong_binary_setoidmor_b f). pose proof (strong_binary_setoidmor_c f). intros x₁ y₁ E1 x₂ y₂ E2. rewrite <-tight_apart in E1. rewrite <-tight_apart in E2. apply tight_apart. intros E3. edestruct (cotransitive E3 (f y₁ x₂)). destruct E1. now apply (strong_extensionality (λ x, f x x₂)). destruct E2. now apply (strong_extensionality (f y₁)). Qed. End morphisms. Section more_morphisms. Context `{StrongSetoid A} `{StrongSetoid B}. Lemma strong_binary_setoid_morphism_commutative {f : A → A → B} `{!Commutative f} `{∀ z, StrongSetoid_Morphism (f z)} : StrongSetoid_BinaryMorphism f. Proof. assert (∀ z, StrongSetoid_Morphism (λ x, f x z)). split; try apply _. intros x y. rewrite !(commutativity _ z). now apply (strong_extensionality (f z)). apply strong_binary_setoid_morphism_both_coordinates. Qed. End more_morphisms. #[global] Instance default_apart `{Equiv A} : Apart A | 20 := (≠). Set Warnings "-unsupported-attributes". (* FIXME: remove when minimal Coq version is enough *) #[global] Typeclasses Opaque default_apart. Set Warnings "+unsupported-attributes". #[global] Instance default_apart_trivial `{Equiv A} : TrivialApart A (Aap:=default_apart). Proof. red. reflexivity. Qed. (* In case we have a decidable setoid, we can construct a strong setoid. Again we do not make this an instance as it will cause loops *) Section dec_setoid. Context `{Setoid A} `{Apart A} `{!TrivialApart A} `{∀ x y, Decision (x = y)}. (* Not Global in order to avoid loops *) Instance ne_apart x y : PropHolds (x ≠ y) → PropHolds (x ≶ y). Proof. rewrite trivial_apart. easy. Qed. Instance dec_strong_setoid: StrongSetoid A. Proof. split; try apply _. firstorder. intros x y. rewrite !trivial_apart. firstorder. intros x y E1 z. rewrite !trivial_apart. destruct (decide (x = z)) as [E2|E2]; [|tauto]. right. intros E3. rewrite trivial_apart in E1. apply E1. now rewrite E2. intros x y. rewrite trivial_apart. split. intros E. now apply stable. firstorder. Qed. End dec_setoid. (* And a similar result for morphisms *) Section dec_setoid_morphisms. Context `{StrongSetoid A} `{!TrivialApart A} `{StrongSetoid B}. Instance dec_strong_morphism (f : A → B) `{!Setoid_Morphism f} : StrongSetoid_Morphism f. Proof. split; try apply _. intros x y E. apply trivial_apart, (setoids.morphism_ne f). now apply apart_ne. Qed. Context `{!TrivialApart B}. Instance dec_strong_injective (f : A → B) `{!Injective f} : StrongInjective f. Proof. pose proof (injective_mor f). split; try apply _. intros x y. rewrite !trivial_apart. now apply (injective_ne f). Qed. Context `{StrongSetoid C}. Instance dec_strong_binary_morphism (f : A → B → C) `{!Proper ((=) ==> (=) ==> (=)) f} : StrongSetoid_BinaryMorphism f. Proof. split; try apply _. intros x₁ y₁ x₂ y₂ E1. case (cotransitive E1 (f x₂ y₁)); rewrite !trivial_apart; intros E2. left. intros E3. destruct (apart_ne _ _ E2). now rewrite E3. right. intros E3. destruct (apart_ne _ _ E2). now rewrite E3. Qed. End dec_setoid_morphisms. math-classes-8.19.0/theory/ua_congruence.v000066400000000000000000000214221460576051100205210ustar00rootroot00000000000000Require Import Coq.Classes.Morphisms Coq.Classes.RelationClasses Coq.Relations.Relation_Definitions Coq.Lists.List MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.interfaces.canonical_names MathClasses.theory.ua_subalgebraT MathClasses.misc.util. Require MathClasses.theory.ua_products. Require MathClasses.theory.categories. (* Remove this *) Local Hint Transparent Equiv : typeclass_instances. Section contents. Context σ `{@Algebra σ v ve vo}. Notation op_type := (op_type (sorts σ)). Notation vv := (ua_products.carrier σ bool (λ _: bool, v)). Instance hint: @Algebra σ vv (ua_products.product_e σ bool (fun _ : bool => v) (fun _ : bool => ve)) _ := @ua_products.product_algebra σ bool (λ _, v) _ _ _. (* Given an equivalence on the algebra's carrier that respects its setoid equality... *) Instance hint' (a: sorts σ): Equiv (ua_products.carrier σ bool (fun _: bool => v) a). Proof. apply products.dep_prod_equiv. intro. apply _. Defined. Context (e: ∀ s, relation (v s)). Section for_nice_e. Context (e_e: ∀ s, Equivalence (e s)) (e_proper: ∀ s, Proper ((=) ==> (=) ==> iff) (e s)). (* We can show that that equivalence lifted to arbitrary operation types still respects the setoid equality: *) Let Q s x := e s (x true) (x false). Let lifted_e := @op_type_equiv (sorts σ) v e. Let lifted_normal := @op_type_equiv (sorts σ) v ve. Instance lifted_e_proper o: Proper ((=) ==> (=) ==> iff) (lifted_e o). Proof with intuition. induction o; simpl. intuition. repeat intro. unfold respectful. split; intros. assert (x x1 = y x1). apply H0... assert (x0 y1 = y0 y1). apply H1... apply (IHo (x x1) (y x1) H4 (x0 y1) (y0 y1) H5)... assert (x x1 = y x1). apply H0... assert (x0 y1 = y0 y1). apply H1... apply (IHo (x x1) (y x1) H4 (x0 y1) (y0 y1) H5)... Qed. (* todo: clean up *) (* With that out of the way, we show that there are two equivalent ways of stating compatibility with the algebra's operations: *) (* 1: the direct way with Algebra; *) Let eAlgebra := @Algebra σ v e _. (* 2: the indirect way of saying that the relation as a set of pairs is a subalgebra in the product algebra: *) Let eSub := @ua_subalgebraT.ClosedSubset σ vv _ _ Q. Lemma eAlgebra_eSub: eAlgebra → eSub. Proof with intuition. intros. constructor. unfold Q. repeat intro. constructor; intro. rewrite <- (H1 true), <- (H1 false)... rewrite (H1 true), (H1 false)... intro o. simpl. unfold algebra_op, ua_products.product_ops, algebra_op. set (f := λ _: bool, vo o). assert (∀ b, Proper (=) (f b)). intro. unfold f. apply algebra_propers. assert (lifted_e _ (f true) (f false)). unfold f. unfold lifted_e. destruct H0. apply algebra_propers. assert (∀ b, Proper (lifted_e (σ o)) (f b))... clearbody f. induction (σ o)... simpl in *... apply IHo0... apply H1... Qed. (* todo: clean up *) Lemma eSub_eAlgebra: eSub → eAlgebra. Proof with intuition. intros [proper closed]. constructor. unfold abstract_algebra.Setoid. apply _. intro o. generalize (closed o). clear closed. (* todo: must be a cleaner way *) unfold algebra_op. simpl. unfold ua_products.product_ops. intro. change (lifted_e _ (algebra_op o) (algebra_op o)). set (f := λ _: bool, algebra_op o) in *. assert (∀ b, lifted_normal _ (f b) (f b)). intros. subst f. simpl. apply algebra_propers... change (lifted_e (σ o) (f true) (f false)). clearbody f. induction (σ o)... simpl in *. repeat intro. unfold respectful in H0. apply (IHo0 (λ b, f b (if b then x else y)))... Qed. (* todo: clean up *) Lemma back_and_forth: iffT eSub eAlgebra. Proof. split; intro; [apply eSub_eAlgebra | apply eAlgebra_eSub]; assumption. Qed. End for_nice_e. (* This justifies the following definition of a congruence: *) Class Congruence: Prop := { congruence_proper:: ∀ s: sorts σ, Proper ((=) ==> (=) ==> iff) (e s) ; congruence_quotient:: Algebra σ v (e:=e) }. End contents. (* A variety for an equational theory none of whose laws have premises is closed under quotients generated by congruences: *) Lemma quotient_variety (et: EquationalTheory) `{InVariety et v} (e': ∀ s, relation (v s)) `{!Congruence et e'} (no_premises: ∀ s, et_laws et s → entailment_premises _ s ≡ nil): InVariety et v (e:=e'). (* Todo: Can this no-premises condition be weakened? Does it occur in this form in the literature? *) Proof. constructor. apply _. intros l law vars. pose proof (variety_laws l law vars) as E. pose proof (no_premises l law). destruct l as [prems [conc ?]]. simpl in *. subst. simpl in *. unfold equiv. rewrite E. pose proof (_: Equivalence (e' conc)). reflexivity. Qed. Section in_domain. Context {A B} {R: Equiv B} (f: A → B). Definition in_domain: Equiv A := λ x y, f x = f y. (* todo: use pointwise thingy *) Global Instance in_domain_equivalence: Equivalence R → Equivalence in_domain. Proof with intuition. constructor; repeat intro; unfold equiv, in_domain in *... Qed. End in_domain. Section first_iso. (* "If A and B are algebras, and f is a homomorphism from A to B, then the equivalence relation Φ on A defined by a~b if and only if f(a)=f(b) is a congruence on A, and the algebra A/Φ is isomorphic to the image of f, which is a subalgebra of B." *) Context `{Algebra σ A} `{Algebra σ B} `{!HomoMorphism σ A B f}. Notation Φ := (λ s, in_domain (f s)). Lemma square o0 x x' y y': Preservation σ A B f x x' → Preservation σ A B f y y' → op_type_equiv (sorts σ) B o0 x' y' → @op_type_equiv (sorts σ) A (λ s: sorts σ, @in_domain (A s) (B s) (e0 s) (f s)) o0 x y. Proof. induction o0. simpl. intros. unfold in_domain. rewrite H3, H4. assumption. simpl in *. repeat intro. unfold in_domain in H6. unfold respectful in H5. simpl in *. pose proof (H3 x0). pose proof (H3 y0). clear H3. pose proof (H4 x0). pose proof (H4 y0). clear H4. apply (IHo0 (x x0) (x' (f _ x0)) (y y0) (y' (f _ y0)) H7 H9). apply H5. assumption. Qed. Instance co: Congruence σ Φ. Proof with intuition. constructor. repeat intro. unfold in_domain. rewrite H3, H4... constructor; intro. unfold abstract_algebra.Setoid. apply _. unfold algebra_op. generalize (preserves σ A B f o). generalize (@algebra_propers σ B _ _ _ o). unfold algebra_op. generalize (H o), (H1 o). induction (σ o); simpl in *; repeat intro. apply _. apply (square _ _ _ _ _ (H4 x) (H4 y))... Qed. Definition image s (b: B s): Type := sigT (λ a, f s a = b). Lemma image_proper: ∀ s (x0 x' : B s), x0 = x' → iffT (image s x0) (image s x'). Proof. intros ??? E. split; intros [v ?]; exists v; rewrite E in *; assumption. Qed. Instance: ClosedSubset image. Proof with intuition. constructor; repeat intro. split; intros [q p]; exists q; rewrite p... generalize (preserves σ A B f o). generalize (@algebra_propers σ B _ _ _ o). unfold algebra_op. generalize (H1 o), (H o). induction (σ o); simpl; intros. exists o1... destruct X. apply (@op_closed_proper σ B _ _ _ image image_proper _ (o1 z) (o1 (f _ x))). apply H3... apply IHo0 with (o2 x)... apply _. Qed. Definition quot_obj: algebras.Object σ := algebras.object σ A (algebra_equiv:=Φ). (* A/Φ *) Definition subobject: algebras.Object σ := algebras.object σ (ua_subalgebraT.carrier image). Program Definition back: subobject ⟶ quot_obj := λ _ X, projT1 (projT2 X). Next Obligation. Proof with try apply _; intuition. repeat constructor... intros [x [i E']] [y [j E'']] E. change (x = y) in E. change (f a i = f a j). rewrite E', E''... unfold ua_subalgebraT.impl. generalize (subset_closed image o). unfold algebra_op. generalize (H o) (H1 o) (preserves σ A B f o) (_: Proper (=) (H o)) (_: Proper (=) (H1 o)). induction (σ o); simpl; intros ? ? ? ? ?. intros [? E]. change (f _ x = f _ o0). rewrite E... intros ? [x [? E]]. apply IHo0... simpl in *. rewrite <- E... Defined. Program Definition forth: quot_obj ⟶ subobject := λ a X, existT _ (f a X) (existT _ X (reflexivity _)). Next Obligation. Proof with try apply _; intuition. repeat constructor... unfold ua_subalgebraT.impl. generalize (subset_closed image o). unfold algebra_op. generalize (H o) (H1 o) (preserves σ A B f o) (_: Proper (=) (H o)) (_: Proper (=) (H1 o)). induction (σ o); simpl... apply IHo0... Qed. Theorem first_iso: categories.iso_arrows back forth. Proof. split. intro. reflexivity. intros ? [? [? ?]]. assumption. Qed. End first_iso. math-classes-8.19.0/theory/ua_homomorphisms.v000066400000000000000000000131671460576051100213040ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra. Section contents. Variable σ: Signature. Notation OpType := (OpType (sorts σ)). Section homo. Context (A B: sorts σ → Type) `{A_equiv : ∀ a, Equiv (A a)} `{B_equiv : ∀ a, Equiv (B a)} `{A_ops : AlgebraOps σ A} `{B_ops : AlgebraOps σ B}. Section with_f. Context (f : ∀ a, A a → B a). Arguments f {a} _. Fixpoint Preservation {n : OpType}: op_type A n → op_type B n → Prop := match n with | ne_list.one d => λ oA oB, f oA = oB | ne_list.cons x y => λ oA oB, ∀ x, Preservation (oA x) (oB (f x)) end. Class HomoMorphism: Prop := { homo_proper:: ∀ a, Setoid_Morphism (@f a) ; preserves: ∀ (o: σ), Preservation (A_ops o) (B_ops o) ; homo_source_algebra: Algebra σ A ; homo_target_algebra: Algebra σ B }. Context `{∀ a, Setoid (A a)} `{∀ b, Setoid (B b)} `{∀ a, Setoid_Morphism (@f a)}. Global Instance Preservation_proper n: Proper (op_type_equiv _ _ _ ==> op_type_equiv _ B n ==> iff) (@Preservation n). (* todo: use (=) in the signature and see why things break *) Proof with auto. induction n; simpl; intros x y E x' y' E'. split; intro F. rewrite <- E, <- E'... rewrite E, E'... split; simpl; intros. eapply IHn; eauto; symmetry; [now apply E | now apply E']. eapply IHn; eauto; [now apply E | now apply E']. Qed. Global Instance Preservation_proper'' n: Proper (eq ==> (=) ==> iff) (@Preservation n). Proof with auto. induction n; simpl; intros x y E x' y' E'. split; intro F. rewrite <- E, <- E'... rewrite E, E'... split; simpl; intros. subst. apply (IHn (y x0) (y x0) eq_refl (y' (f x0)) (x' (f x0)) ). symmetry. apply E'. reflexivity. apply H2. subst. apply (IHn (y x0) (y x0) eq_refl (y' (f x0)) (x' (f x0)) ). symmetry. apply E'. reflexivity. apply H2. Qed. (* todo: evil, get rid of *) End with_f. Lemma Preservation_proper' (f g: ∀ a, A a → B a) `{∀ a, Setoid (A a)} `{∀ b, Setoid (B b)} `{∀ a, Setoid_Morphism (@f a)}: (∀ a (x: A a), f a x = g a x) → ∀ (n: OpType) x y, Proper (=) x → Proper (=) y → @Preservation f n x y → @Preservation g n x y. Proof. induction n. simpl. intros ? ? ? ? E. rewrite <-E. symmetry. intuition. simpl. intros a b E1 E2 E3 z. apply IHn. apply E1. reflexivity. apply E2. reflexivity. assert (b (g _ z) = b (f _ z)) as E4. apply E2. symmetry. apply H2. now apply (Preservation_proper'' f n (a z) (a z) eq_refl _ _ E4). Qed. Lemma HomoMorphism_Proper: Proper ((λ f g, ∀ a x, f a x = g a x) ==> iff) HomoMorphism. (* todo: use pointwise_thingy *) Proof with try apply _; intuition. intros x y E1. constructor; intros [? ? ? ?]; simpl in *. repeat constructor... intros ? ? E2. rewrite <-2!E1. rewrite E2... apply (Preservation_proper' x y E1 (σ o) (A_ops o) (B_ops o))... repeat constructor... intros ? ? E2. rewrite 2!E1. rewrite E2... assert (∀ (a : sorts σ) (x0 : A a), y a x0 = x a x0) as E2. symmetry. apply E1. apply (Preservation_proper' y x E2 (σ o) (A_ops o) (B_ops o))... Qed. End homo. Global Instance id_homomorphism A `{∀ a, Equiv (A a)} {ao: AlgebraOps σ A} `{!Algebra σ A}: HomoMorphism _ _ (λ _, id). Proof with try apply _; intuition. constructor; intros... generalize (ao o). induction (σ o); simpl... reflexivity. Qed. Global Instance compose_homomorphisms A B C f g `{∀ a, Equiv (A a)} `{∀ a, Equiv (B a)} `{∀ a, Equiv (C a)} {ao: AlgebraOps σ A} {bo: AlgebraOps σ B} {co: AlgebraOps σ C} {gh: HomoMorphism A B g} {fh: HomoMorphism B C f}: HomoMorphism A C (λ a, f a ∘ g a). Proof with try apply _; auto. pose proof (homo_source_algebra _ _ g). pose proof (homo_target_algebra _ _ g). pose proof (homo_target_algebra _ _ f). constructor; intros... generalize (ao o) (bo o) (co o) (preserves _ _ g o) (preserves _ _ f o). induction (σ o); simpl; intros; unfold compose. rewrite H5... apply (IHo0 _ (o2 (g _ x)))... Qed. Lemma invert_homomorphism A B f `{∀ a, Equiv (A a)} `{∀ a, Equiv (B a)} {ao: AlgebraOps σ A} {bo: AlgebraOps σ B} {fh: HomoMorphism A B f} `{inv: ∀ a, Inverse (f a)}: (∀ a, Bijective (f a)) → HomoMorphism A B f → HomoMorphism B A inv. Proof with try assumption; try apply _. intros ? [? ? ? ?]. constructor... intro. generalize (ao o) (bo o) (preserves _ _ f o) (algebra_propers o: Proper (=) (ao o)) (algebra_propers o: Proper (=) (bo o)). induction (σ o); simpl. intros. apply (injective (f t)). pose proof (surjective (f t) o1 o1 (reflexivity o1)). transitivity o1... symmetry... intros P Q R S T x. apply IHo0. eapply Preservation_proper''; eauto; intros; try apply _. symmetry. now apply T, (surjective (f t) x x). apply S. reflexivity. apply T. reflexivity. Qed. (* Instance eval_morphism `{Algebra σ} A {V} (v: Vars σ A V): HomoMorphism (Term0 σ V) A (λ _ => eval σ v). Proof. constructor; repeat intro; try apply _. unfold AlgebraOps_instance_0. generalize (algebra_propers o: eval v (Op V o) = H o). generalize (Op V o) (H o). induction (operation_type σ o); simpl; repeat intro. assumption. apply IHo0. simpl. apply H1. destruct H0. reflexivity. Qed. *) End contents. math-classes-8.19.0/theory/ua_mapped_operations.v000066400000000000000000000031411460576051100221000ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra. Section contents. Variable Sorts: Set. Section map_op. (* Given maps between two realizations of the sorts, there are maps between the corresponding op_types*) Context {A B: Sorts → Type} `{∀ a, Equiv (A a)} `{∀ a, Equiv (B a)} (ab: ∀ a, A a → B a) (ba: ∀ a, B a → A a) `{∀ a, Proper ((=) ==> (=)) (ab a)} `{∀ a, Proper ((=) ==> (=)) (ba a)}. Fixpoint map_op {o: OpType Sorts}: op_type A o → op_type B o := match o return op_type A o → op_type B o with | ne_list.one u => ab u | ne_list.cons _ _ => λ x y, map_op (x (ba _ y)) end. Global Instance map_op_proper o: Proper ((=) ==> (=)) (@map_op o). Proof. unfold equiv. induction o; simpl; firstorder. Qed. (* todo: can't we make this nameless? *) End map_op. (* If the maps between the sorts are eachother's inverse, then so are the two generated op_type maps: *) Context {A B: Sorts → Type} {e: ∀ a, Equiv (B a)} `{∀ b, Equivalence (e b)} (ab: ∀ a, A a → B a) (ba: ∀ a, B a → A a). Arguments ab [a] _. Arguments ba [a] _. Context `(iso: ∀ a (x: B a), ab (ba x) = x). Lemma map_iso o (x: op_type B o) (xproper: Proper (=) x): map_op ab ba (map_op ba ab x) = x. Proof with auto; try reflexivity. induction o; simpl; intros... intros x0 y H0. change (map_op ab ba (map_op ba ab (x (ab (ba x0)))) = x y). transitivity (x (ab (ba x0))). apply IHo, xproper... apply xproper. rewrite iso, H0... Qed. End contents. math-classes-8.19.0/theory/ua_packed.v000066400000000000000000000065061460576051100176260ustar00rootroot00000000000000Require Import MathClasses.interfaces.canonical_names MathClasses.interfaces.universal_algebra Coq.Program.Program. Section packed. Context (σ: Signature) {V: Type}. Inductive Applied: sorts σ → Type := | AppliedOp o: Arguments (σ o) → Applied (result _ (σ o)) | AppliedVar: V → ∀ s, Applied s with Arguments: OpType (sorts σ) → Type := | NoMoreArguments s: Arguments (ne_list.one s) | MoreArguments o' o: Applied o' → Arguments o → Arguments (ne_list.cons o' o). Definition head_arg x y (a: Arguments (ne_list.cons x y)): Applied x := match a with | MoreArguments k l m n => m end. Definition tail_args x y (a: Arguments (ne_list.cons x y)): Arguments y := match a with | MoreArguments k l m n => n end. Context (P: ∀ {s}, Applied s → Type). Arguments P {s} _. Section forallArgs. Fixpoint forallArgs {o} (a: Arguments o): Type := match a with | NoMoreArguments _ => True | MoreArguments _ _ z v => prod (P z) (forallArgs v) end. Definition PofSplit {o}: Arguments o → Type := match o with | ne_list.one _ => λ _, unit | ne_list.cons x y => λ a, prod (P (head_arg x y a)) (forallArgs (tail_args _ _ a)) end. Definition forallSplit `(a: Arguments (ne_list.cons y z)): forallArgs a → prod (P (head_arg y z a)) (forallArgs (tail_args _ _ a)). Proof. intro. change (PofSplit a). destruct a; simpl. exact tt. assumption. Defined. (* todo: write as term *) End forallArgs. Context (Pop: ∀ o x, forallArgs x → P (AppliedOp o x)) (Pvar: ∀ v s, P (AppliedVar v s)). Fixpoint better_Applied_rect {o} (a: Applied o): P a := match a with | AppliedOp x y => Pop x y (better_args y) | AppliedVar x v => Pvar x v end with better_args {o} (a: Arguments o): forallArgs a := match a with | NoMoreArguments _ => I | MoreArguments _ _ x y => (better_Applied_rect x, better_args y) end. End packed. (* Conversion /from/ packed representation: *) Fixpoint curry {σ} {V} {o} (a: Applied σ o): Term σ V (ne_list.one o) := match a in (Applied _ s) (*return (Term σ V (ne_list.one s))*) with | AppliedOp _ op y => apply_args y (app_tree σ (Op σ V op)) | AppliedVar _ v x => Var σ V v x end with apply_args {σ} {V} {o} (a: @Arguments σ V o): op_type (Term0 σ V) o → Term0 σ V (result _ o) := match a with | NoMoreArguments _ y => id | MoreArguments _ x y u q => λ z, apply_args q (z (curry u)) end. (* Conversion /to/ packed representation: *) Fixpoint decode `(t: Term σ V o): Arguments σ o → Applied σ (result _ o) := match t with | Var _ _ x y => λ z, AppliedVar σ x y | Op _ _ x => AppliedOp σ x | App _ _ x y z w => λ p, decode z (MoreArguments σ y x (decode w (NoMoreArguments σ _)) p) end. (* Back and forth: *) Definition curry_decode `(t: Term σ V o) (a: Arguments σ o): curry (decode t a) ≡ apply_args a (app_tree σ t). Proof with simpl in *; try reflexivity. induction t... dependent destruction a... (* DANGER: uses JMeq_eq *) rewrite IHt1... rewrite IHt2... Qed. (* Nullary specializations: *) Definition decode0 `(t: Term0 σ V s): Applied σ s := decode t (NoMoreArguments σ _). Definition curry_decode0 `(t: Term0 σ V o): curry (decode0 t) ≡ t. Proof. apply (curry_decode t). Qed. math-classes-8.19.0/theory/ua_products.v000066400000000000000000000175751460576051100202520ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.theory.categories MathClasses.categories.varieties. Require MathClasses.theory.setoids. Section algebras. Context (sig: Signature) (I: Type) (carriers: I → sorts sig → Type) `(∀ i s, Equiv (carriers i s)) `{∀ i, AlgebraOps sig (carriers i)} `{∀ i, Algebra sig (carriers i)}. Definition carrier: sorts sig → Type := λ sort, ∀ i: I, carriers i sort. Fixpoint rec_impl o: (∀ i, op_type (carriers i) o) → op_type carrier o := match o return (∀ i, op_type (carriers i) o) → op_type carrier o with | ne_list.one _ => id | ne_list.cons _ g => λ X X0, rec_impl g (λ i, X i (X0 i)) end. Instance u (s: sorts sig): Equiv (∀ i : I, carriers i s) := products.dep_prod_equiv _ _. Instance rec_impl_proper: ∀ o, Proper (@products.dep_prod_equiv I _ (fun _ => op_type_equiv _ _ _) ==> (=)) (rec_impl o). Proof with auto. induction o; simpl. repeat intro... intros ? ? Y x0 y0 ?. apply IHo. intros. intro. apply Y. change (x0 i = y0 i)... Qed. Global Instance product_ops: AlgebraOps sig carrier := λ o, rec_impl (sig o) (λ i, algebra_op o). Instance: ∀ o: sig, Proper (=) (algebra_op o: op_type carrier (sig o)). Proof. intro. apply rec_impl_proper. intro. apply (algebra_propers _). Qed. Instance product_e sort: Equiv (carrier sort) := (=). (* hint; shouldn't be needed *) Global Instance product_algebra: Algebra sig carrier := {}. Lemma preservation i o: Preservation sig carrier (carriers i) (λ _ v, v i) (algebra_op o) (algebra_op o). unfold product_ops, algebra_op. set (select_op := λ i0 : I, H0 i0 o). replace (H0 i o) with (select_op i) by reflexivity. clearbody select_op. revert select_op. induction (operation_type sig o). simpl. reflexivity. intros. intro. apply IHo0. Qed. Lemma algebra_projection_morphisms i: @HomoMorphism sig carrier (carriers i) _ _ _ _ (λ a v, v i). Proof. constructor; try apply _. intro. rapply (@products.dep_prod_morphism I (λ i, carriers i a) (λ i, _: Equiv (carriers i a))). intro. apply _. apply preservation. Qed. End algebras. Section varieties. Context (et: EquationalTheory) (I: Type) (carriers: I → sorts et → Type) `(∀ i s, Equiv (carriers i s)) `(∀ i, AlgebraOps et (carriers i)) `(∀ i, InVariety et (carriers i)). Notation carrier := (carrier et I carriers). Instance carrier_e : forall s, Equiv _ := product_e et I carriers _. Fixpoint nqe {t}: op_type carrier t → (∀ i, op_type (carriers i) t) → Prop := match t with | ne_list.one _ => λ f g, ∀ i, f i = g i | ne_list.cons _ _ => λ f g, ∀ tuple, nqe (f tuple) (λ i, g i (tuple i)) end. (* todo: rename *) Instance nqe_proper t: Proper ((=) ==> (λ x y, ∀ i, x i = y i) ==> iff) (@nqe t). Proof with auto; try reflexivity. induction t; simpl; intros ? ? P ? ? E. intuition. rewrite <- (P i), <- E... rewrite (P i), E... split; intros U tuple; apply (IHt (x tuple) (y tuple) (P tuple tuple (reflexivity _)) (λ u, x0 u (tuple u)) (λ u, y0 u (tuple u)))... intro. apply E... intro. apply E... Qed. Lemma nqe_always {t} (term: Term _ nat t) vars: nqe (eval et vars term) (λ i, eval et (λ sort n, vars sort n i) term). Proof. induction term; simpl in *. reflexivity. set (k := (λ i : I, eval et (λ (sort: sorts et) (n : nat), vars sort n i) term1 (eval et vars term2 i))). cut (nqe (eval et vars term1 (eval et vars term2)) k). set (p := λ i : I, eval et (λ (sort : sorts et) (n : nat), vars sort n i) term1 (eval et (λ (sort : sorts et) (n : nat), vars sort n i) term2)). assert (op_type_equiv (sorts et) carrier t (eval et vars term1 (eval et vars term2)) (eval et vars term1 (eval et vars term2))). apply sig_type_refl. intro. apply _. apply eval_proper; try apply _. reflexivity. reflexivity. apply (nqe_proper t (eval et vars term1 (eval et vars term2)) (eval et vars term1 (eval et vars term2)) H2 k p). subst p k. simpl. intro. pose proof (eval_proper et _ (λ (sort : sorts et) (n : nat), vars sort n i) (λ (sort : sorts et) (n : nat), vars sort n i) (reflexivity _) term1 term1 eq_refl). apply H3. apply IHterm2. apply IHterm1. change (nqe (rec_impl et _ _ (et o) (λ i : I, @algebra_op _ (carriers i) _ o)) (λ i : I, @algebra_op _ (carriers i) _ o)). generalize (λ i: I, @algebra_op et (carriers i) _ o). induction (et o); simpl. reflexivity. intros. apply IHo0. Qed. Lemma component_equality_to_product t (A A': op_type carrier t) (B B': ∀ i, op_type (carriers i) t): (∀ i, B i = B' i) → nqe A B → nqe A' B' → A = A'. Proof with auto. induction t; simpl in *; intros BE ABE ABE'. intro. rewrite ABE, ABE'... intros x y U. apply (IHt (A x) (A' y) (λ i, B i (x i)) (λ i, B' i (y i)))... intros. apply BE, U. Qed. Lemma laws_hold s: et_laws et s → ∀ vars, eval_stmt et vars s. Proof. intros. pose proof (λ i, variety_laws s H2 (λ sort n, vars sort n i)). clear H2. assert (∀ i : I, eval_stmt (et_sig et) (λ (sort : sorts (et_sig et)) (n : nat), vars sort n i) (entailment_as_conjunctive_statement (et_sig et) s)). intros. rewrite <- boring_eval_entailment. apply (H3 i). clear H3. rewrite boring_eval_entailment. destruct s. simpl in *. destruct entailment_conclusion. simpl in *. destruct i. simpl in *. intro. apply component_equality_to_product with (λ i, eval et (λ sort n, vars sort n i) t) (λ i, eval et (λ sort n, vars sort n i) t0). intro. apply H2. clear H2. simpl. induction entailment_premises... simpl in *. intuition. simpl. rewrite <- (nqe_always (fst (projT2 a)) vars i). rewrite <- (nqe_always (snd (projT2 a)) vars i). intuition. apply H3. apply IHentailment_premises. apply H3. apply nqe_always. apply nqe_always. Qed. (* todo: clean up! also, we shouldn't have to go through boring.. *) Global Instance product_variety: InVariety et carrier (e:=carrier_e). Proof. constructor. apply product_algebra. intro i. apply _. apply laws_hold. Qed. End varieties. Require MathClasses.categories.varieties. Section categorical. Context (et: EquationalTheory). Global Instance: Producer (varieties.Object et) := λ I carriers, {| varieties.variety_carriers := λ s, ∀ i, carriers i s ; varieties.variety_proof := product_variety et I _ _ _ (fun H => varieties.variety_proof et (carriers H)) |}. (* todo: clean up *) Section for_a_given_c. Context (I: Type) (c: I → varieties.Object et). Global Program Instance: ElimProduct c (product c) := λ i _ c, c i. Next Obligation. apply (@algebra_projection_morphisms et I c (λ x, @varieties.variety_equiv et (c x)) (λ x, varieties.variety_ops et (c x)) ). intro. apply _. Qed. Global Program Instance: IntroProduct c (product c) := λ H h a X i, h i a X. Next Obligation. Proof with intuition. repeat constructor; try apply _. intros ?? E ?. destruct h. simpl. rewrite E... intro o. pose proof (λ i, @preserves _ _ _ _ _ _ _ _ (proj2_sig (h i)) o) as H0. unfold product_ops, algebra_op. set (o0 := λ i, varieties.variety_ops et (c i) o). set (o1 := varieties.variety_ops et H o) in *. change (∀i : I, Preservation et H (c i) (` (h i)) o1 (o0 i)) in H0. clearbody o0 o1. revert o0 o1 H0. induction (et o); simpl... apply (@product_algebra et I c)... Qed. Global Instance: Product c (product c). Proof. split; repeat intro. reflexivity. symmetry. simpl. apply H. Qed. End for_a_given_c. Global Instance: HasProducts (varieties.Object et) := {}. End categorical. (* Todo: Lots of cleanup. *) math-classes-8.19.0/theory/ua_subalgebra.v000066400000000000000000000062031460576051100205000ustar00rootroot00000000000000Require Import Coq.Classes.RelationClasses MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.interfaces.canonical_names MathClasses.theory.categories MathClasses.interfaces.abstract_algebra. Require MathClasses.categories.algebras MathClasses.theory.forget_algebra. Section subalgebras. Context `{Algebra sign A} (P: ∀ s, A s → Prop). (* We begin by describing what it means for P to be a proper closed subset: *) Fixpoint op_closed {o: OpType (sorts sign)}: op_type A o → Prop := match o with | ne_list.one x => P x | ne_list.cons _ _ => λ d, ∀ z, P _ z → op_closed (d z) end. Global Instance op_closed_proper: ∀ `{∀ s, Proper ((=) ==> iff) (P s)} o, Proper ((=) ==> iff) (@op_closed o). Proof with intuition. induction o; simpl; intros x y E. rewrite E... split; intros. apply (IHo (x z))... apply E... apply (IHo _ (y z))... apply E... Qed. Class ClosedSubset: Prop := { subset_proper:: ∀ s, Proper ((=) ==> iff) (P s) ; subset_closed: ∀ o, op_closed (algebra_op o) }. (* Now suppose P is closed in this way. *) Context `{ClosedSubset}. (* Our new algebra's elements are just those for which P holds: *) Definition carrier s := sig (P s). Hint Extern 4 => progress unfold carrier: typeclass_instances. (* We can implement closed operations in the new algebra: *) Program Fixpoint close_op {d}: ∀ (o: op_type A d), op_closed o → op_type carrier d := match d with | ne_list.one _ => λ o c, exist _ o (c) | ne_list.cons _ _ => λ o c X, close_op (o X) (c X (proj2_sig X)) end. Global Instance impl: AlgebraOps sign carrier := λ o, close_op (algebra_op o) (subset_closed o). (* By showing that these ops are proper, we get our new algebra: *) Definition close_op_proper d (o0 o1: op_type A d) (P': op_closed o0) (Q: op_closed o1): o0 = o1 → close_op o0 P' = close_op o1 Q. Proof with intuition. induction d; simpl in *... intros [x p] [y q] U. apply (IHd _ _ (P' x p) (Q y q)), H2... Qed. Global Instance subalgebra: Algebra sign carrier. Proof. constructor. apply _. intro. apply close_op_proper, algebra_propers. Qed. (* And we have the obvious projection morphism: *) Definition proj s := @proj1_sig (A s) (P s). Global Instance: HomoMorphism sign carrier A proj. Proof with try apply _. constructor... constructor... firstorder. intro. unfold impl, algebra_op. generalize (subset_closed o). unfold algebra_op. generalize (H o). induction (sign o); simpl; intuition. Qed. (* Which is mono because proj is injective. *) Instance: `{Injective (proj i)}. Proof. split. firstorder. apply (@homo_proper sign carrier A (fun s : sorts sign => @sig_equiv (A s) (e s) (P s)) _ _ _ _). apply _. Qed. Global Instance: Mono (algebras.arrow _ proj). Proof. apply forget_algebra.mono. apply categories.product.mono. intro. apply setoids.mono. simpl. apply _. Qed. (* this really should be completely automatic. *) End subalgebras. #[global] Hint Extern 10 (Equiv (carrier _ _)) => apply @sig_equiv : typeclass_instances. math-classes-8.19.0/theory/ua_subalgebraT.v000066400000000000000000000052551460576051100206320ustar00rootroot00000000000000(* This is an almost verbatim copy of ua_subalgebra, but with predicates in Type instead of Prop (and with carrier sigT instead of sig). Hopefully one day Coq's universe polymorphism will permit a merge of sig and sigT, at which point we may try to merge ua_subalgebra and ua_subalgebraT as well. *) Require Import Coq.Classes.RelationClasses MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.theory.categories MathClasses.interfaces.abstract_algebra. Require MathClasses.categories.algebras. Section subalgebras. Context `{Algebra sign A} (P: ∀ s, A s → Type). (* We begin by describing what it means for P to be a proper closed subset: *) Fixpoint op_closed {o: OpType (sorts sign)}: op_type A o → Type := match o with | ne_list.one x => P x | ne_list.cons x y => λ d, ∀ z, P _ z → op_closed (d z) end. Definition op_closed_proper: ∀ (Pproper: ∀ s x x', x = x' → iffT (P s x) (P s x')) o, ∀ x x', x = x' → iffT (@op_closed o x) (@op_closed o x'). Proof with intuition. induction o; simpl; intros x y E. intuition. split; intros... apply (IHo (x z))... apply E... apply (IHo (y z))... symmetry. apply E... Qed. Class ClosedSubset: Type := { subset_proper: ∀ s x x', x = x' → iffT (P s x) (P s x') ; subset_closed: ∀ o, op_closed (algebra_op o) }. (* Now suppose P is closed in this way. *) Context `{ClosedSubset}. (* Our new algebra's elements are just those for which P holds: *) Definition carrier s := sigT (P s). Hint Extern 4 => progress unfold carrier: typeclass_instances. (* We can implement closed operations in the new algebra: *) Fixpoint close_op {d}: ∀ (o: op_type A d), op_closed o → op_type carrier d := match d with | ne_list.one _ => λ o c, existT _ o (c) | ne_list.cons _ _ => λ o c X, close_op (o (projT1 X)) (c (projT1 X) (projT2 X)) end. Global Instance impl: AlgebraOps sign carrier := λ o, close_op (algebra_op o) (subset_closed o). (* By showing that these ops are proper, we get our new algebra: *) Instance: ∀ d, Equiv (op_type carrier d). intro. apply op_type_equiv. intro. apply _. Defined. Definition close_op_proper d (o0 o1: op_type A d) (P': op_closed o0) (Q: op_closed o1): o0 = o1 → close_op o0 P' = close_op o1 Q. Proof with intuition. induction d; simpl in *... intros [x p] [y q] U. apply (IHd _ _ (P' x p) (Q y q)), H2... Qed. Global Instance subalgebra: Algebra sign carrier. Proof. constructor. apply _. intro. apply close_op_proper, algebra_propers. Qed. End subalgebras. #[global] Hint Extern 4 => progress unfold carrier: typeclass_instances. math-classes-8.19.0/theory/ua_subvariety.v000066400000000000000000000061471460576051100205750ustar00rootroot00000000000000Require Import Coq.Classes.RelationClasses Coq.Classes.Morphisms Coq.Program.Program MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.interfaces.canonical_names MathClasses.theory.ua_subalgebra. (* In theory/ua_subalgebra.v we defined closed proper subsets and showed that they yield subalgebras. We now expand on this result and show that they also yield subvarieties (by showing that the laws still hold in the subalgebra). *) Section contents. Context `{InVariety et A} `{@ClosedSubset et A _ _ P}. (* todo: why so ugly? *) Definition Pvars (vars: Vars et (carrier P) nat): Vars et A nat := λ s n, ` (vars s n). (* To prove that the laws still hold in the subalgebra, we first prove that evaluation in it is the same as evaluation in the original: *) Program Fixpoint heq {o}: op_type (carrier P) o → op_type A o → Prop := match o with | ne_list.one _ => λ a b, `a = b | ne_list.cons _ _ => λ a b, ∀ u, heq (a u) (b u) end. Instance heq_proper {o}: Proper ((=) ==> (=) ==> iff) (@heq o). Proof with intuition. intros x y U x0 y0 K. induction o; simpl in *. destruct x, y. change (x = x1) in U. simpl in *. split; intro. transitivity x... transitivity x0... transitivity x1... transitivity y0... assert (∀ u, x u = y u). intros. apply U... split; repeat intro. apply -> (IHo (x u) (y u) (H1 u) (x0 (proj1_sig u)))... apply K... apply <- (IHo (x u) (y u) (H1 u) (x0 (proj1_sig u)))... apply K... Qed. Lemma heq_eval vars {o} (t: T et o): heq (eval et vars t) (eval et (Pvars vars) t). Proof with intuition. induction t; simpl... unfold Pvars... simpl in IHt1. generalize (IHt1 (eval et vars t3)). clear IHt1. apply heq_proper. pose proof (@eval_proper et (carrier P) _ _ _ nat (ne_list.cons y t1)). apply H1; try intro... pose proof (@eval_proper et A _ _ _ nat (ne_list.cons y t1)). apply H1... unfold heq in IHt2. (* todo: this wasn't needed in a previous Coq version *) rewrite IHt2. apply (@eval_proper et A _ _ _ nat (ne_list.one y))... unfold impl, algebra_op. generalize (subset_closed P o). unfold algebra_op. generalize (AlgebraOps0 o). intros. induction (et o); simpl in *... Qed. Lemma heq_eval_const vars {o} (t: T et (ne_list.one o)): ` (eval et vars t) = eval et (Pvars vars) t. Proof. apply (heq_eval vars t). Qed. (* todo: this specialization wasn't needed in a previous Coq version *) Lemma laws s: et_laws et s → ∀ vars: ∀ a, nat → carrier P a, eval_stmt et vars s. Proof with intuition. intros. generalize (@variety_laws et A _ _ _ s H1 (Pvars vars)). clear H1. destruct s as [x [? [t t0]]]. induction x as [| [x1 [t1 t2]]]; simpl in *; intros. unfold equiv, sig_equiv. rewrite (heq_eval_const vars t). rewrite (heq_eval_const vars t0)... apply IHx, H1. rewrite <- (heq_eval_const vars t1). rewrite <- (heq_eval_const vars t2)... Qed. (* Which gives us our variety: *) Global Instance: InVariety et (carrier P) := { variety_laws := laws }. End contents. math-classes-8.19.0/theory/ua_term_monad.v000066400000000000000000000071421460576051100205210ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.interfaces.monads. Section contents. Context (operation: Set) (operation_type: operation → OpType unit). Let sign := Build_Signature unit operation operation_type. (* The monad's type constructor: *) Definition M (T: Type): Type := Term sign T (ne_list.one tt). (* We first define equality for arbitrary arities, and then instantiate for constants. *) Section equality. Context `{Setoid A}. Fixpoint geneq {s s'} (x: Term sign A s) (y: Term sign A s'): Prop := match x, y with | Var _ _ v _, Var _ _ w _ => v = w | App _ _ _ z t t', App _ _ _ z' t'' t''' => geneq t t'' ∧ geneq t' t''' | Op _ _ o, Op _ _ o' => o ≡ o' | _, _ => False end. Lemma geneq_sym s (x: Term sign A s): ∀ s' (y: Term sign A s'), geneq x y → geneq y x. Proof with intuition. induction x. destruct y... simpl. symmetry... destruct y0... simpl in *... destruct y... simpl in *... Qed. Lemma geneq_trans s (x: Term sign A s): ∀ s' (y: Term sign A s') s'' (z: Term sign A s''), geneq x y → geneq y z → geneq x z. Proof with simpl in *; intuition. induction x; simpl. destruct y, z... destruct y0... destruct z... eauto. eauto. destruct y, z... transitivity o0... Qed. Global Instance gen_equiv: Equiv (M A) := geneq. Global Instance: Setoid (M A). Proof. split. intros x. unfold M, equiv, gen_equiv in *. induction x; simpl; intuition. repeat intro. now apply geneq_sym. repeat intro. now apply geneq_trans with _ y. Qed. End equality. (* For bind, we do the same: *) Definition gen_bind_aux {A B: Type} (f: A → M B): ∀ {s}, Term sign A s → Term sign B s := fix F {s} (t: Term sign A s): Term sign B s := match t with | Var _ _ v tt => f v | App _ _ o z x y => App _ _ _ z (F x) (F y) | Op _ _ o => Op _ _ o end. Arguments gen_bind_aux {A B} _ {s} _. Instance gen_bind: MonadBind M := λ _ _ f z, gen_bind_aux f z. Instance: ∀ `{Equiv A} `{Equiv B}, Proper (((=) ==> (=)) ==> (=) ==> (=)) (@bind M _ A B). Proof with intuition. intros A H1 B H2 x0 y0 E' x y E. revert x y E. change (∀ x y : M A, geneq x y → geneq (gen_bind_aux x0 x) (gen_bind_aux y0 y)). cut (∀ s (x: Term sign A s) s' (y: Term sign A s'), geneq x y → geneq (gen_bind_aux x0 x) (gen_bind_aux y0 y))... revert s' y H. induction x. destruct y... simpl in *. destruct a, a1. apply E'... simpl in *. destruct y... destruct y... simpl in *... destruct y... Qed. (* return: *) Instance gen_ret: MonadReturn M := λ _ x, Var sign _ x tt. Instance: ∀ `{Equiv A}, Proper ((=) ==> (=)) (@ret M _ A). Proof. repeat intro. assumption. Qed. (* What remains are the laws: *) Instance: Monad M. Proof with intuition. constructor; intros; try apply _. (* law 1 *) now apply setoids.ext_equiv_refl. (* law 2 *) intros m n E. rewrite <-E. clear E n. unfold M in m. change (geneq (gen_bind_aux (λ x : A, Var sign A x tt) m) m). induction m; simpl... destruct a... simpl... (* law 3 *) intros m n E. rewrite E. clear E m. unfold bind, gen_bind, equiv, gen_equiv. revert n. assert (∀ o (m : Term sign A o), geneq (gen_bind_aux f (gen_bind_aux g m)) (gen_bind_aux (λ x : A, gen_bind_aux f (g x)) m)) as H. induction m; simpl... destruct a. change (gen_bind_aux f (g v) = gen_bind_aux f (g v))... now apply H. Qed. End contents. math-classes-8.19.0/theory/ua_transference.v000066400000000000000000000065431460576051100210570ustar00rootroot00000000000000Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.interfaces.canonical_names MathClasses.theory.categories MathClasses.theory.ua_mapped_operations. Require MathClasses.categories.varieties. Section contents. Context (et: EquationalTheory) `{InVariety et A} `{InVariety et B} `{!HomoMorphism et A B ab} `{!HomoMorphism et B A ba} (i: iso_arrows (varieties.arrow et ab) (varieties.arrow et ba)). Arguments ab {a} _. Arguments ba {a} _. Let ab_proper a: Proper ((=) ==> (=)) (@ab a). Proof. apply _. Qed. Let ba_proper a: Proper ((=) ==> (=)) (@ba a). Proof. apply _. Qed. Let epA: ∀ V n, Proper ((=) ==> eq ==> (=)) (@eval _ A _ V n) := _. Let epB: ∀ V n, Proper ((=) ==> eq ==> (=)) (@eval _ B _ V n) := _. (* hints. shouldn't be necessary *) Let ab_ba: ∀ b (a: B b), ab (ba a) = a := proj1 i. Let ba_ab: ∀ b (a: A b), ba (ab a) = a := proj2 i. Program Lemma transfer_eval n (t: Term et nat n) (v: Vars et B nat): eval et (A:=A) (λ _ i, ba (v _ i)) t = map_op _ (@ba) (@ab) (eval et v t). Proof with auto; try reflexivity. induction t; simpl in *; intros... set (eval et (λ (a: sorts et) (i : nat), ba (v a i)) t2). eapply transitivity. apply (@epA nat (ne_list.cons y t1) (λ a i, ba (v a i)) (λ a i, ba (v a i)) (reflexivity _) t2 t2 (reflexivity _) : Proper ((=) ==> op_type_equiv (sorts et) A t1)%signature o). now apply IHt2. subst o. rewrite (IHt1 (ba (eval et v t3)) (ba (eval et v t3)))... apply (@map_op_proper (sorts et) B A _ _ _ _ _ _). unfold compose in *. rapply (epB _ _ v v (reflexivity _) t2 t2 (reflexivity _)). apply ab_ba. generalize (@algebra_propers _ A _ _ _ o) (@algebra_propers _ B _ _ _ o). generalize (@preserves et A B _ _ _ _ (@ab) _ o). fold (@algebra_op et A _ o) (@algebra_op et B _ o). generalize (@algebra_op et A _ o), (@algebra_op et B _ o). induction (et o); simpl; repeat intro. rewrite <- ba_ab, H1... apply IHo0. eapply Preservation_proper''; eauto; intros; try apply _. symmetry. apply H3, ab_proper, H4. apply H2... apply H3... Qed. (* todo: make [reflexivity] work as a hint. further cleanup. *) Program Lemma transfer_eval' n (t: Term et nat n) (v: Vars et B nat): map_op _ (@ab) (@ba) (eval et (A:=A) (λ _ i, ba (v _ i)) t) = eval et v t. Proof with auto. intros. pose proof (@map_op_proper (sorts et) A B _ _ _ _ _ _). rewrite transfer_eval. apply (@map_iso _ A B _ _ (@ab) (@ba) ab_ba). apply _. Qed. Program Lemma transfer_statement_and_vars (s: Statement et) (v: Vars et B nat): eval_stmt et v s ↔ eval_stmt et (A:=A) (λ _ i, ba (v _ i)) s. Proof with auto; reflexivity. intros. induction s; simpl; intuition. rewrite transfer_eval. symmetry. rewrite transfer_eval. symmetry. apply (map_op_proper _ _ _)... rewrite <- transfer_eval'. symmetry. rewrite <- transfer_eval'. symmetry. apply (map_op_proper _ _ _)... Qed. Theorem transfer_statement (s: Statement et): (∀ v, eval_stmt et (A:=A) v s) → (∀ v, eval_stmt et (A:=B) v s). Proof with intuition auto. intros ? v. assert (v = (λ _ i, ab (ba (v _ i)))) as P. destruct i. intros a a0. symmetry... rewrite P, transfer_statement_and_vars... Qed. End contents. math-classes-8.19.0/varieties/000077500000000000000000000000001460576051100161755ustar00rootroot00000000000000math-classes-8.19.0/varieties/abgroup.v000066400000000000000000000142471460576051100200330ustar00rootroot00000000000000(* To be imported qualified. *) Require MathClasses.categories.varieties MathClasses.categories.product MathClasses.theory.forget_algebra MathClasses.theory.forget_variety MathClasses.theory.groups. Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.misc.workaround_tactics categories.categories. Inductive op := mult | inv | one. Definition sig: Signature := single_sorted_signature (λ o, match o with one => O | inv => 1%nat | mult => 2%nat end). Section laws. Global Instance: SgOp (Term0 sig nat tt) := fun x => App sig _ _ _ (App sig _ _ _ (Op sig nat mult) x). Global Instance: MonUnit (Term0 sig nat tt) := Op sig nat one. Global Instance: Negate (Term0 sig nat tt) := fun x => App sig _ _ _ (Op sig nat inv) x. Local Notation x := (Var sig nat 0%nat tt). Local Notation y := (Var sig nat 1%nat tt). Local Notation z := (Var sig nat 2%nat tt). Import notations. Inductive Laws: EqEntailment sig → Prop := | e_mult_assoc: Laws (x & (y & z) === (x & y) & z) | e_mult_1_l: Laws (mon_unit & x === x) | e_mult_1_r: Laws (x & mon_unit === x) | e_mult_inv_l: Laws (-x & x === mon_unit) | e_mult_inv_r: Laws (x & -x === mon_unit) | e_mult_commute: Laws (x & y === y & x). End laws. Definition theory: EquationalTheory := Build_EquationalTheory sig Laws. Definition Object := varieties.Object theory. #[global] Hint Extern 4 => match goal with [ |- ∀ _:unit, _ ] => intros [] end : typeclass_instances. Program Definition forget: Object → setoids.Object := product.project (λ _, setoids.Object) tt ∘ forget_algebra.object theory ∘ forget_variety.forget theory. (* todo: too ugly *) (* Now follow a series of encoding/decoding functions to convert between the specialized Monoid/Monoid_Morphism type classes and the universal Algebra/InVariety/HomoMorphism type classes instantiated with the above signature and theory. *) #[global] Instance encode_operations A `{!SgOp A} `{Negate A} `{!MonUnit A}: AlgebraOps sig (λ _, A) := λ o, match o with mult => sg_op | inv => negate | one => mon_unit: A end. Section decode_operations. Context `{AlgebraOps theory A}. Global Instance: MonUnit (A tt) := algebra_op one. Global Instance: SgOp (A tt) := algebra_op mult. Global Instance: Negate (A tt) := algebra_op inv. End decode_operations. Section encode_variety_and_ops. Context A `{AbGroup A}. Global Instance encode_algebra_and_ops: Algebra sig _. Proof. constructor. intro. apply _. intros []; simpl; try apply _; unfold Proper; reflexivity. Qed. Hint Resolve associativity left_identity right_identity left_inverse right_inverse commutativity. Global Instance encode_variety_and_ops: InVariety theory (λ _, A). Proof. constructor. apply _. intros ? [] ?; simpl; unfold algebra_op; simpl; eauto. apply associativity. apply left_identity. apply right_identity. eapply left_inverse. apply right_inverse. apply commutativity. Qed. Definition object: Object := varieties.object theory (λ _, A). End encode_variety_and_ops. Lemma encode_algebra_only `{!AlgebraOps theory A} `{∀ u, Equiv (A u)} `{!AbGroup (A tt)}: Algebra theory A . Proof. constructor; intros []; simpl in *; try apply _. Qed. Global Instance decode_variety_and_ops `{InVariety theory A}: AbGroup (A tt). Proof with simpl; auto. pose proof (λ law lawgood x y z, variety_laws law lawgood (λ s n, match s with tt => match n with 0 => x | 1 => y | _ => z end end)) as laws. repeat constructor; try apply _; hnf; intros. - apply_simplified (laws _ e_mult_assoc). - apply_simplified (algebra_propers mult)... - apply_simplified (laws _ e_mult_1_l)... - apply_simplified (laws _ e_mult_1_r)... - apply_simplified (algebra_propers inv)... - apply_simplified (laws _ e_mult_inv_l)... - apply_simplified (laws _ e_mult_inv_r)... - apply_simplified (laws _ e_mult_commute)... Qed. Lemma encode_morphism_only `{AlgebraOps theory A} `{∀ u, Equiv (A u)} `{AlgebraOps theory B} `{∀ u, Equiv (B u)} (f: ∀ u, A u → B u) `{!AbGroup (A tt)} `{!AbGroup (B tt)} `{!Monoid_Morphism (f tt)}: HomoMorphism sig A B f. Proof. constructor. * intros []. apply _. * intros []; simpl. + apply preserves_sg_op. + apply groups.preserves_negate. + apply (@preserves_mon_unit (A tt) (B tt) _ _ _ _ _ _ (f tt) _). * eapply encode_algebra_only. * eapply encode_algebra_only. Qed. Lemma encode_morphism_and_ops `{AbGroup A} `{AbGroup B} `{!Monoid_Morphism (f: A→B)}: @HomoMorphism sig (λ _, A) (λ _, B) _ _ ( _) ( _) (λ _, f). Proof. intros. apply encode_morphism_only; assumption. Qed. Lemma decode_morphism_and_ops `{InVariety theory x} `{InVariety theory y} `{!HomoMorphism theory x y f}: Monoid_Morphism (f tt). Proof. pose proof (homo_proper theory x y f tt). constructor; try apply _. constructor; try apply _. apply (preserves theory x y f mult). apply (preserves theory x y f one). Qed. (* Finally, we use these encoding/decoding functions to specialize some universal results: *) Section specialized. Context (A B C: Type) `{!MonUnit A} `{!SgOp A} `{!Negate A} `{!Equiv A} `{!MonUnit B} `{!SgOp B} `{!Negate B} `{!Equiv B} `{!MonUnit C} `{!SgOp C} `{!Negate C} `{!Equiv C} (f: A → B) (g: B → C). Global Instance id_morphism `{!AbGroup A}: Monoid_Morphism (id:A->A). Proof. repeat (constructor; try apply _); reflexivity. Qed. Global Instance compose_morphisms `{!AbGroup A, !AbGroup B, !AbGroup C} `{!Monoid_Morphism f} `{!Monoid_Morphism g}: Monoid_Morphism (g ∘ f). Proof. pose proof (encode_morphism_and_ops (f:=f)) as P. pose proof (encode_morphism_and_ops (f:=g)) as Q. pose proof (@compose_homomorphisms theory _ _ _ _ _ _ _ _ _ _ _ P Q). apply (@decode_morphism_and_ops _ _ _ _ _ _ _ _ _ H). Qed. Global Instance: ∀ `{!AbGroup A, !AbGroup B} `{H: !Monoid_Morphism (f: A->B)} `{!Inverse f}, Bijective f → Monoid_Morphism (inverse f). Proof. intros. pose proof (encode_morphism_and_ops (f:=f)) as P. pose proof (@invert_homomorphism theory _ _ _ _ _ _ _ _ _ _ P) as Q. destruct H. apply (@decode_morphism_and_ops _ _ _ _ _ _ _ _ _ Q). Qed. End specialized. math-classes-8.19.0/varieties/closed_terms.v000066400000000000000000000210521460576051100210470ustar00rootroot00000000000000Require Import Coq.Classes.RelationClasses Coq.Relations.Relation_Definitions Coq.Lists.List Coq.Classes.Morphisms MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.interfaces.abstract_algebra MathClasses.interfaces.canonical_names MathClasses.theory.categories MathClasses.categories.varieties. Section contents. Variable et: EquationalTheory. (* The initial object will consists of arity-0 terms with a congruence incorporating the variety's laws. For this we simply take universal_algebra's Term type, but exclude variables by taking False as the variable index type: *) Let ClosedTerm := (Term et False). Let ClosedTerm0 a := ClosedTerm (ne_list.one a). (* Operations are implemented as App-tree builders, so that [o a b] yields [App (App (Op o) a) b]. *) Fixpoint app_tree {o}: ClosedTerm o → op_type ClosedTerm0 o := match o with | ne_list.one _ => id | ne_list.cons _ _ => λ x y, app_tree (App _ _ _ _ x y) end. Instance: AlgebraOps et ClosedTerm0 := λ x, app_tree (Op _ _ x). (* We define term equivalence on all operation types: *) Inductive e: ∀ o, Equiv (ClosedTerm o) := | e_refl o: Reflexive (e o) | e_trans o: Transitive (e o) | e_sym o: Symmetric (e o) | e_sub o h: Proper ((=) ==> (=) ==> (=)) (App _ _ h o) | e_law (s: EqEntailment et): et_laws et s → ∀ (v: Vars et ClosedTerm0 nat), (∀ x, In x (entailment_premises _ s) → eval et v (fst (projT2 x)) = eval et v (snd (projT2 x))) → eval et v (fst (projT2 (entailment_conclusion _ s))) = eval et v (snd (projT2 (entailment_conclusion _ s))). Existing Instance e. Existing Instance e_refl. Existing Instance e_sym. Existing Instance e_trans. (* .. and then take the specialization at arity 0 for Term0: *) Instance: ∀ a, Equiv (ClosedTerm0 a) := λ a, e (ne_list.one a). Instance: ∀ a, Setoid (ClosedTerm0 a). Proof. split; apply _. Qed. (* While this fancy congruence is the one we'll use to make our initial object a setoid, in our proofs we will also need to talk about extensionally equal closed term builders (i.e. terms of type [op_type ClosedTerm0 a] for some a), where we use /Leibniz/ equality on closed terms: *) Let structural_eq a: relation _ := @op_type_equiv (sorts et) ClosedTerm0 (λ _, eq) a. Instance structural_eq_refl a: Reflexive (structural_eq a). Proof. induction a; repeat intro. reflexivity. subst. apply IHa. Qed. (* The implementation is proper: *) Instance app_tree_proper: ∀ o, Proper ((=) ==> (=)) (@app_tree o). Proof with auto. induction o; repeat intro... apply IHo, e_sub... Qed. Instance: Algebra et ClosedTerm0. Proof. constructor. intro. apply _. intro. apply app_tree_proper. reflexivity. Qed. (* hm, this isn't "the" (closed) term algebra, because we used equivalence modulo an equational theory *) (* Better still, the laws hold: *) Lemma laws_hold s (L: et_laws et s): ∀ vars, eval_stmt _ vars s. Proof with simpl in *; intuition. intros. rewrite boring_eval_entailment. destruct s. simpl in *. intros. apply (e_law _ L vars). clear L. induction entailment_premises... subst... Qed. (* Hence, we have a variety: *) Global Instance: InVariety et ClosedTerm0. Proof. constructor. apply _. intros. apply laws_hold. assumption. Qed. (* And hence, an object in the category: *) Definition the_object: varieties.Object et := varieties.object et ClosedTerm0. (* Interestingly, this object is initial, which we prove in the remainder of this file. *) (* To show its initiality, we begin by constructing arrows to arbitrary other objects: *) Section for_another_object. Variable other: varieties.Object et. (* Computationally, the arrow simply evaluates closed terms in the other model. For induction purposes, we first define this for arbitrary op_types: *) Definition eval_in_other {o}: ClosedTerm o → op_type other o := @eval et other _ False o (no_vars _ other). Definition morph a: the_object a → other a := eval_in_other. (* Given an assignment mapping variables to closed terms of arity 0, we can now evaluate open terms in the other model in two ways: by first closing it and then evaluating the closed term using eval_in_other, or by evaluating it directly but with variable lookup implemented in terms of eval_in_other. We now show that these two ways yield the same result, thanks to proper-ness of the algebra's operations: *) Lemma subst_eval o V (v: Vars _ ClosedTerm0 _) (t: Term _ V o): @eval _ other _ _ _ (λ x y, eval_in_other (v x y)) t = eval_in_other (close _ v t). Proof. induction t; simpl. reflexivity. apply IHt1. auto. apply (@algebra_propers et other _ _ _ o). Qed. (* todo: rename *) (* On the side of the_object, evaluating a term of arity 0 is the same as closing it: *) Lemma eval_is_close V x v (t: Term0 et V x): eval et v t ≡ close _ v t. Proof with auto; try reflexivity. pattern x, t. apply applications_rect; simpl... intro. cut (@equiv _ (structural_eq _) (app_tree (close _ v (Op et _ o))) (eval et v (Op et _ o)))... generalize (Op et V o). induction (et o); simpl... intros ? H ? E. apply IHo0. simpl. rewrite E. apply H... Qed. (* And with those two somewhat subtle lemmas, we show that eval_in_other is a setoid morphism: *) Instance prep_proper {o}: Proper ((=) ==> (=)) (@eval_in_other o). Proof with intuition. intros x y H. induction H; simpl... induction x; simpl... apply IHx1... apply (@algebra_propers et other _ _ _ o). apply IHe... unfold Vars in v. pose proof (@variety_laws et other _ _ _ s H (λ a n, eval_in_other (v a n))) as Q. clear H. destruct s. rewrite boring_eval_entailment in Q. simpl in *. do 2 rewrite eval_is_close. do 2 rewrite <- subst_eval. apply Q. clear Q. induction entailment_premises; simpl... do 2 rewrite subst_eval. do 2 rewrite <- eval_is_close... Qed. Instance: ∀ a, Setoid_Morphism (@eval_in_other (ne_list.one a)). Proof. constructor; simpl; try apply _. Qed. (* Furthermore, we can show preservation of operations, giving us a homomorphism (and an arrow): *) Instance: @HomoMorphism et ClosedTerm0 other _ (varieties.variety_equiv et other) _ _ (λ _, eval_in_other). Proof with intuition. constructor; try apply _. intro. change (Preservation et ClosedTerm0 other (λ _, eval_in_other) (app_tree (Op _ _ o)) (varieties.variety_ops _ other o)). generalize (algebra_propers o : eval_in_other (Op _ _ o) = varieties.variety_ops _ other o). generalize (Op _ False o) (varieties.variety_ops et other o). induction (et o)... simpl. intro. apply IHo0, H. apply reflexivity. (* todo: shouldn't have to say [apply] here. file bug *) Qed. Program Definition the_arrow: the_object ⟶ other := λ _, eval_in_other. (* All that remains is to show that this arrow is unique: *) Theorem arrow_unique: ∀ y, the_arrow = y. Proof with auto; try intuition. intros [x h] b a. simpl in *. pattern b, a. apply applications_rect... pose proof (@preserves et ClosedTerm0 other _ _ _ _ x h o). change (Preservation et ClosedTerm0 other x (app_tree (Op _ _ o)) (@eval_in_other _ (Op _ _ o))) in H. revert H. generalize (Op _ False o). induction (et o); simpl... apply IHo0. simpl in *. assert (app_tree (App _ _ o0 t t0 v) = app_tree (App _ _ o0 t t0 v)). apply app_tree_proper... apply (@Preservation_proper et ClosedTerm0 other _ _ x _ _ _ o0 (app_tree (App _ _ o0 t t0 v)) (app_tree (App _ _ o0 t t0 v)) H1 (eval_in_other t0 (eval_in_other v)) (eval_in_other t0 (x t v)))... pose proof (@prep_proper _ t0 t0 (reflexivity _))... (* todo: why doesn't rewrite work here? *) Qed. (* todo: needs further cleanup *) End for_another_object. Hint Extern 4 (InitialArrow the_object) => exact the_arrow: typeclass_instances. Instance: Initial the_object. Proof. intro. apply arrow_unique. Qed. (* Todo: Show decidability of equality (likely quite tricky). Without that, we cannot use any of this to get a canonical model of the naturals/integers, because such models must be decidable. *) (* (Tom): This is possible in the general case, since that would involve solving the word problem for groups. In particular, a group presentation determines a variety, the initial object of which is the group so presented. *) End contents. math-classes-8.19.0/varieties/empty.v000066400000000000000000000015651460576051100175310ustar00rootroot00000000000000Require MathClasses.theory.rings MathClasses.categories.varieties. Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra. Definition op := False. Definition sig: Signature := Build_Signature False False (False_rect _). Definition Laws (_: EqEntailment sig): Prop := False. Definition theory: EquationalTheory := Build_EquationalTheory sig Laws. #[local] Definition carriers := False_rect _: sorts sig → Type. #[global] Instance: `{Equiv (carriers a)}. Proof. intros []. Qed. #[global] Instance implementation: AlgebraOps sig carriers := λ o, False_rect _ o. Global Instance: Algebra sig _. Proof. constructor; intuition. Qed. #[global] Instance variety: InVariety theory carriers. Proof. constructor; intuition. Qed. Definition Object := varieties.Object theory. Definition object: Object := varieties.object theory carriers. math-classes-8.19.0/varieties/groups.v000066400000000000000000000113611460576051100177050ustar00rootroot00000000000000(* To be imported qualified. *) Require MathClasses.categories.varieties MathClasses.categories.product MathClasses.theory.forget_algebra MathClasses.theory.forget_variety MathClasses.theory.groups. Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.misc.workaround_tactics categories.categories. Inductive op := mult | inv | one. Definition sig: Signature := single_sorted_signature (λ o, match o with one => O | inv => 1%nat | mult => 2%nat end). Section laws. Global Instance: SgOp (Term0 sig nat tt) := λ x, App sig _ _ _ (App sig _ _ _ (Op sig nat mult) x). Global Instance: MonUnit (Term0 sig nat tt) := Op sig nat one. Global Instance: Negate (Term0 sig nat tt) := λ x, App sig _ _ _ (Op sig nat inv) x. Local Notation x := (Var sig nat 0%nat tt). Local Notation y := (Var sig nat 1%nat tt). Local Notation z := (Var sig nat 2%nat tt). Import notations. Inductive Laws: EqEntailment sig → Prop := | e_mult_assoc: Laws (x & (y & z) === (x & y) & z) | e_mult_1_l: Laws (mon_unit & x === x) | e_mult_1_r: Laws (x & mon_unit === x) | e_recip_l: Laws (-x & x === mon_unit) | e_recip_r: Laws (x & -x === mon_unit). End laws. Definition theory: EquationalTheory := Build_EquationalTheory sig Laws. Definition Object := varieties.Object theory. Definition forget: Object → setoids.Object := @product.project unit (λ _, setoids.Object) (λ _, _: Arrows setoids.Object) _ (λ _, _: CatId setoids.Object) (λ _, _: CatComp setoids.Object) (λ _, _: Category setoids.Object) tt ∘ forget_algebra.object theory ∘ forget_variety.forget theory. (* todo: too ugly *) (* Now follow a series of encoding/decoding functions to convert between the specialized Monoid/Monoid_Morphism type classes and the universal Algebra/InVariety/HomoMorphism type classes instantiated with the above signature and theory. *) #[global] Instance encode_operations A `{!SgOp A} `(Negate A) `{!MonUnit A}: AlgebraOps sig (λ _, A) := λ o, match o with mult => (&) | inv => (-) | one => mon_unit: A end. Section decode_operations. Context `{AlgebraOps theory A}. Global Instance: MonUnit (A tt) := algebra_op one. Global Instance: SgOp (A tt) := algebra_op mult. Global Instance: Negate (A tt) := algebra_op inv. End decode_operations. Section encode_variety_and_ops. Context A `{Group A}. Global Instance encode_algebra_and_ops: Algebra sig _. Proof. constructor. intro. apply _. intro o. destruct o; simpl; try apply _; unfold Proper; reflexivity. Qed. Global Instance encode_variety_and_ops: InVariety theory (λ _, A). Proof. constructor. apply _. intros ? [] ?; simpl; unfold algebra_op; simpl. apply associativity. apply left_identity. apply right_identity. eapply left_inverse. eapply right_inverse. Qed. Definition object: Object := varieties.object theory (λ _, A). End encode_variety_and_ops. Lemma encode_algebra_only `{!AlgebraOps theory A} `{∀ u, Equiv (A u)} `{!Group (A tt)}: Algebra theory A . Proof. constructor; intros []; simpl in *; try apply _. Qed. Global Instance decode_variety_and_ops `{InVariety theory A}: Group (A tt). Proof with simpl; auto. pose proof (λ law lawgood x y z, variety_laws law lawgood (λ s n, match s with tt => match n with 0 => x | 1 => y | _ => z end end)) as laws. repeat constructor; try apply _. intro. apply_simplified (laws _ e_mult_assoc). apply_simplified (algebra_propers mult)... intro. apply_simplified (laws _ e_mult_1_l)... intro. apply_simplified (laws _ e_mult_1_r)... apply_simplified (algebra_propers inv). intro. apply_simplified (laws _ e_recip_l)... intro. apply_simplified (laws _ e_recip_r)... Qed. Lemma encode_morphism_only `{AlgebraOps theory A} `{∀ u, Equiv (A u)} `{AlgebraOps theory B} `{∀ u, Equiv (B u)} (f: ∀ u, A u → B u) `{!Group (A tt)} `{!Group (B tt)} `{!Monoid_Morphism (f tt)}: HomoMorphism sig A B f. Proof. pose proof (monmor_a (f:=f tt)). pose proof (monmor_b (f:=f tt)). constructor. intros []. apply _. intros []; simpl. apply preserves_sg_op. apply groups.preserves_negate. rapply preserves_mon_unit. eapply encode_algebra_only. eapply encode_algebra_only. Qed. Lemma encode_morphism_and_ops `{Group A} `{Group B} `{!Monoid_Morphism (f: A → B)}: @HomoMorphism sig (λ _, A) (λ _, B) _ _ ( _) ( _) (λ _, f). Proof. intros. apply encode_morphism_only; assumption. Qed. Lemma decode_morphism_and_ops `{InVariety theory x} `{InVariety theory y} `{!HomoMorphism theory x y f}: Monoid_Morphism (f tt). Proof. constructor; try apply _. constructor; try apply _. apply (preserves theory x y f mult). apply (preserves theory x y f one). Qed. math-classes-8.19.0/varieties/monoids.v000066400000000000000000000132201460576051100200320ustar00rootroot00000000000000(* To be imported qualified. *) Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.misc.workaround_tactics MathClasses.categories.categories. Require MathClasses.categories.varieties MathClasses.categories.product MathClasses.theory.forget_algebra MathClasses.theory.forget_variety. Inductive op := mult | one. Definition sig: Signature := single_sorted_signature (λ o, match o with one => O | mult => 2%nat end). Section laws. Global Instance: SgOp (Term0 sig nat tt) := fun x => App sig _ _ _ (App sig _ _ _ (Op sig nat mult) x). Global Instance: MonUnit (Term0 sig nat tt) := Op sig nat one. Local Notation x := (Var sig nat 0%nat tt). Local Notation y := (Var sig nat 1%nat tt). Local Notation z := (Var sig nat 2%nat tt). Import notations. Inductive Laws: EqEntailment sig → Prop := | e_mult_assoc: Laws (x & (y & z) === (x & y) & z) | e_mult_1_l: Laws (mon_unit & x === x) | e_mult_1_r: Laws (x & mon_unit === x). End laws. Definition theory: EquationalTheory := Build_EquationalTheory sig Laws. Definition Object := varieties.Object theory. Local Hint Extern 3 => progress simpl : typeclass_instances. Definition forget: Object → setoids.Object := @product.project unit (λ _, setoids.Object) (λ _, _) _ (λ _, _) (λ _, _) (λ _, _) tt ∘ forget_algebra.object theory ∘ forget_variety.forget theory. (* todo: too ugly *) (* Now follow a series of encoding/decoding functions to convert between the specialized Monoid/Monoid_Morphism type classes and the universal Algebra/InVariety/HomoMorphism type classes instantiated with the above signature and theory. *) #[global] Instance encode_operations A `{!SgOp A} `{!MonUnit A}: AlgebraOps sig (λ _, A) := λ o, match o with mult => (&) | one => mon_unit: A end. Section decode_operations. Context `{AlgebraOps theory A}. Global Instance: MonUnit (A tt) := algebra_op one. Global Instance: SgOp (A tt) := algebra_op mult. End decode_operations. Section encode_variety_and_ops. Context A `{Monoid A}. Global Instance encode_algebra_and_ops: Algebra sig _. Proof. constructor. intro. apply _. intro o. destruct o; simpl; try apply _; unfold Proper; reflexivity. Qed. Global Instance encode_variety_and_ops: InVariety theory (λ _, A) | 10. Proof. constructor. apply _. intros ? [] ?; simpl; unfold algebra_op; simpl. apply associativity. apply left_identity. apply right_identity. Qed. Definition object: Object := varieties.object theory (λ _, A). End encode_variety_and_ops. Lemma encode_algebra_only `{!AlgebraOps theory A} `{∀ u, Equiv (A u)} `{!Monoid (A tt)}: Algebra theory A . Proof. constructor; intros []; apply _. Qed. Global Instance decode_variety_and_ops `{InVariety theory A}: Monoid (A tt) | 10. Proof with simpl; auto. pose proof (λ law lawgood x y z, variety_laws law lawgood (λ s n, match s with tt => match n with 0 => x | 1 => y | _ => z end end)) as laws. constructor. constructor. apply _. intro. apply_simplified (laws _ e_mult_assoc). apply_simplified (algebra_propers mult)... intro. apply_simplified (laws _ e_mult_1_l)... intro. apply_simplified (laws _ e_mult_1_r)... Qed. Lemma encode_morphism_only `{AlgebraOps theory A} `{∀ u, Equiv (A u)} `{AlgebraOps theory B} `{∀ u, Equiv (B u)} (f: ∀ u, A u → B u) `{!Monoid_Morphism (f tt)}: HomoMorphism sig A B f. Proof. pose proof (monmor_a (f:=f tt)). pose proof (monmor_b (f:=f tt)). constructor. intros []. apply _. intros []; simpl. apply preserves_sg_op. apply (@preserves_mon_unit (A tt) (B tt) _ _ _ _ _ _ (f tt)). apply _. apply encode_algebra_only. apply encode_algebra_only. Qed. Lemma encode_morphism_and_ops `{Monoid_Morphism A B f}: @HomoMorphism sig (λ _, A) (λ _, B) _ _ ( _) ( _) (λ _, f). Proof. intros. apply encode_morphism_only. assumption. Qed. Lemma decode_morphism_and_ops `{InVariety theory x} `{InVariety theory y} `{!HomoMorphism theory x y f}: Monoid_Morphism (f tt). Proof. constructor; try apply _. constructor; try apply _. apply (preserves theory x y f mult). apply (preserves theory x y f one). Qed. #[global] Instance id_monoid_morphism `{Monoid A}: Monoid_Morphism (@id A). Proof. repeat (split; try apply _); easy. Qed. (* Finally, we use these encoding/decoding functions to specialize some universal results: *) Section specialized. Context `{Equiv A}`{MonUnit A} `{SgOp A} `{Equiv B} `{MonUnit B} `{SgOp B} `{Equiv C} `{MonUnit C} `{SgOp C} (f : A → B) (g : B → C). Instance compose_monoid_morphism: Monoid_Morphism f → Monoid_Morphism g → Monoid_Morphism (g ∘ f). Proof. intros. pose proof (encode_morphism_and_ops (f:=f)) as P. pose proof (encode_morphism_and_ops (f:=g)) as Q. pose proof (@compose_homomorphisms theory _ _ _ _ _ _ _ _ _ _ _ P Q) as PP. pose proof (monmor_a (f:=f)). pose proof (monmor_b (f:=f)). pose proof (monmor_b (f:=g)). apply (@decode_morphism_and_ops _ _ _ _ _ _ _ _ _ PP). Qed. Lemma invert_monoid_morphism: ∀ `{!Inverse f}, Bijective f → Monoid_Morphism f → Monoid_Morphism (f⁻¹). Proof. intros. pose proof (encode_morphism_and_ops (f:=f)) as P. pose proof (@invert_homomorphism theory _ _ _ _ _ _ _ _ _ _ P) as Q. pose proof (monmor_a (f:=f)). pose proof (monmor_b (f:=f)). apply (@decode_morphism_and_ops _ _ _ _ _ _ _ _ _ Q). Qed. End specialized. #[global] Hint Extern 4 (Monoid_Morphism (_ ∘ _)) => class_apply @compose_monoid_morphism : typeclass_instances. #[global] Hint Extern 4 (Monoid_Morphism (_⁻¹)) => class_apply @invert_monoid_morphism : typeclass_instances. math-classes-8.19.0/varieties/open_terms.v000066400000000000000000000070621460576051100205440ustar00rootroot00000000000000Require Import Coq.Classes.RelationClasses Coq.Relations.Relation_Definitions Coq.Lists.List Coq.Classes.Morphisms MathClasses.interfaces.universal_algebra MathClasses.interfaces.abstract_algebra MathClasses.interfaces.canonical_names MathClasses.theory.categories MathClasses.categories.varieties. Section contents. Context (operation: Set) (operation_type: operation → OpType unit). Let sig := Build_Signature unit operation operation_type. Context (laws: EqEntailment sig → Prop). Let et := Build_EquationalTheory sig laws. Context `{Setoid A}. Notation OpenTerm := (Term et A). (* Using a Let brings out Coq bug #2299. *) Definition OpenTerm0 a := OpenTerm (ne_list.one a). (* Operations are implemented as App-tree builders, so that [o a b] yields [App (App (Op o) a) b]. *) Fixpoint app_tree {o}: OpenTerm o → op_type OpenTerm0 o := match o with | ne_list.one _ => id | ne_list.cons _ _ => λ x y, app_tree (App _ _ _ _ x y) end. Instance: AlgebraOps et OpenTerm0 := λ x, app_tree (Op _ _ x). (* We define term equivalence on all operation types: *) Inductive ee: ∀ o, Equiv (OpenTerm o) := | e_vars a a': a = a' → Var et A a tt = Var et A a' tt | e_refl o: Reflexive (ee o) | e_trans o: Transitive (ee o) | e_sym o: Symmetric (ee o) | e_sub o h: Proper ((=) ==> (=) ==> (=)) (App _ _ h o) | e_law (s: EqEntailment et): et_laws et s → ∀ (v: Vars et OpenTerm0 nat), (∀ x, In x (entailment_premises _ s) → eval et v (fst (projT2 x)) = eval et v (snd (projT2 x))) → eval et v (fst (projT2 (entailment_conclusion _ s))) = eval et v (snd (projT2 (entailment_conclusion _ s))). Existing Instance ee. Existing Instance e_refl. Existing Instance e_sym. Existing Instance e_trans. (* .. and then take the specialization at arity 0 for Term0: *) Instance: ∀ a, Equiv (OpenTerm0 a) := λ a, ee (ne_list.one a). Instance: ∀ a, Setoid (OpenTerm0 a). Proof. split; apply _. Qed. (* While this fancy congruence is the one we'll use to make our initial object a setoid, in our proofs we will also need to talk about extensionally equal closed term builders (i.e. terms of type [op_type ClosedTerm0 a] for some a), where we use /Leibniz/ equality on closed terms: *) Let structural_eq a: relation _ := @op_type_equiv unit OpenTerm0 (λ _, eq) a. Instance structural_eq_refl a: Reflexive (structural_eq a). Proof. induction a; repeat intro. reflexivity. subst. apply IHa. Qed. (* The implementation is proper: *) Instance app_tree_proper: ∀ o, Proper ((=) ==> (=)) (@app_tree o). Proof with auto. induction o; repeat intro... apply IHo, e_sub... Qed. Instance: Algebra et OpenTerm0. Proof. constructor. intro. apply _. intro. apply app_tree_proper. reflexivity. Qed. (* Better still, the laws hold: *) Lemma laws_hold s (L: et_laws et s): ∀ vars, eval_stmt _ vars s. Proof with simpl in *; intuition. intros. rewrite boring_eval_entailment. destruct s. simpl in *. intros. apply (e_law _ L vars). clear L. induction entailment_premises... subst... Qed. (* Hence, we have an object in the variety: *) Global Instance: InVariety et OpenTerm0. Proof. constructor. apply _. intros. apply laws_hold. assumption. Qed. Definition the_object: varieties.Object et := varieties.object et OpenTerm0. End contents. (* Todo: - show freedom; - tackle multisorted case (likely requires modification of Term definition because it will no longer facilitate the trick of using the carrier as the variable indexing set). *) math-classes-8.19.0/varieties/rings.v000066400000000000000000000124361460576051100175140ustar00rootroot00000000000000(* To be imported qualified. *) Require MathClasses.categories.varieties MathClasses.theory.rings. Require Import Coq.setoid_ring.Ring MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.misc.workaround_tactics. Inductive op := plus | mult | zero | one | negate. Definition sig: Signature := single_sorted_signature (λ o, match o with zero | one => O | negate => 1%nat | plus | mult => 2%nat end). Section laws. Global Instance: Plus (Term0 sig nat tt) := λ x, App sig _ _ _ (App sig _ _ _ (Op sig _ plus) x). Global Instance: Mult (Term0 sig nat tt) := λ x, App sig _ _ _ (App sig _ _ _ (Op sig _ mult) x). Global Instance: Zero (Term0 sig nat tt) := Op sig _ zero. Global Instance: One (Term0 sig nat tt) := Op sig _ one. Global Instance: Negate (Term0 sig nat tt) := App sig _ _ _ (Op sig _ negate). Local Notation x := (Var sig nat 0%nat tt). Local Notation y := (Var sig nat 1%nat tt). Local Notation z := (Var sig nat 2%nat tt). Import notations. Inductive Laws: EqEntailment sig → Prop := |e_plus_assoc: Laws (x + (y + z) === (x + y) + z) |e_plus_comm: Laws (x + y === y + x) |e_plus_0_l: Laws (0 + x === x) |e_mult_assoc: Laws (x * (y * z) === (x * y) * z) |e_mult_comm: Laws (x * y === y * x) |e_mult_1_l: Laws (1 * x === x) |e_mult_0_l: Laws (0 * x === 0) |e_distr: Laws (x * (y + z) === x * y + x * z) |e_distr_l: Laws ((x + y) * z === x * z + y * z) |e_plus_negate_r: Laws (x + - x === 0) |e_plus_negate_l: Laws (- x + x === 0). End laws. Definition theory: EquationalTheory := Build_EquationalTheory sig Laws. Definition Object := varieties.Object theory. (* Now follow a series of encoding/decoding functions to convert between the specialized Ring/Ring_Morphism type classes and the universal Algebra/InVariety/HomoMorphism type classes instantiated with the above signature and theory. *) Section decode_operations. Context `{AlgebraOps theory A}. Global Instance: Plus (A tt) := algebra_op plus. Global Instance: Mult (A tt) := algebra_op mult. Global Instance: Zero (A tt) := algebra_op zero. Global Instance: One (A tt) := algebra_op one. Global Instance: Negate (A tt) := algebra_op negate. End decode_operations. Section encode_with_ops. Context A `{Ring A}. Global Instance encode_operations: AlgebraOps sig (λ _, A) := λ o, match o with plus => (+) | mult => (.*.) | zero => 0:A | one => 1:A | negate => (-) end. Global Instance encode_algebra_and_ops: Algebra sig _. Proof. constructor. intro. apply _. intro o. destruct o; simpl; try apply _; apply reflexivity. Qed. Add Ring A: (rings.stdlib_ring_theory A). Global Instance encode_variety_and_ops: InVariety theory (λ _, A). Proof. constructor. apply _. intros ? [] ?; simpl; unfold algebra_op; simpl; ring. Qed. Definition object: Object := varieties.object theory (λ _, A). End encode_with_ops. Lemma encode_algebra_only `{!AlgebraOps theory A} `{∀ u, Equiv (A u)} `{!Ring (A tt)}: Algebra sig A . Proof. constructor; intros []; simpl in *; try apply _. Qed. #[global] Instance decode_variety_and_ops `{InVariety theory A}: Ring (A tt). Proof with simpl; auto. pose proof (λ law lawgood x y z, variety_laws law lawgood (λ s n, match s with tt => match n with 0 => x | 1 => y | _ => z end end)) as laws. repeat (constructor; try apply _); repeat intro. apply_simplified (laws _ e_plus_assoc). apply_simplified (algebra_propers plus)... apply_simplified (laws _ e_plus_0_l)... transitivity (algebra_op plus (algebra_op zero) x). apply_simplified (laws _ e_plus_comm)... apply_simplified (laws _ e_plus_0_l)... apply_simplified (algebra_propers negate)... apply_simplified (laws _ e_plus_negate_l)... apply_simplified (laws _ e_plus_negate_r)... apply_simplified (laws _ e_plus_comm)... apply_simplified (laws _ e_mult_assoc)... apply_simplified (algebra_propers mult)... apply_simplified (laws _ e_mult_1_l)... transitivity (algebra_op mult (algebra_op one) x). apply_simplified (laws _ e_mult_comm)... apply_simplified (laws _ e_mult_1_l)... apply_simplified (laws _ e_mult_comm)... apply_simplified (laws _ e_distr)... Qed. Lemma encode_morphism_only `{AlgebraOps theory A} `{∀ u, Equiv (A u)} `{AlgebraOps theory B} `{∀ u, Equiv (B u)} (f: ∀ u, A u → B u) `{!Ring (A tt)} `{!Ring (B tt)} `{!SemiRing_Morphism (f tt)}: HomoMorphism sig A B f. Proof. constructor. intros []. apply _. intros []; simpl. apply rings.preserves_plus. apply rings.preserves_mult. apply rings.preserves_0. apply rings.preserves_1. apply rings.preserves_negate. apply encode_algebra_only. apply encode_algebra_only. Qed. Lemma encode_morphism_and_ops `{Ring A} `{Ring B} {f : A → B} `{!SemiRing_Morphism f}: @HomoMorphism sig (λ _, A) (λ _, B) _ _ _ _ (λ _, f). Proof. intros. apply (encode_morphism_only _). Qed. Lemma decode_morphism_and_ops `{InVariety theory x} `{InVariety theory y} `{!HomoMorphism theory x y f}: SemiRing_Morphism (f tt). Proof. pose proof (preserves theory x y f) as P. repeat (constructor; try apply _) ; [ apply (P plus) | apply (P zero) | apply (P mult) | apply (P one) ]. Qed. math-classes-8.19.0/varieties/semigroups.v000066400000000000000000000120361460576051100205630ustar00rootroot00000000000000(* To be imported qualified. *) Require MathClasses.categories.varieties MathClasses.categories.product MathClasses.theory.forget_algebra MathClasses.theory.forget_variety. Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.misc.workaround_tactics categories.categories. Inductive op := mult. Definition sig: Signature := single_sorted_signature (λ o, match o with mult => 2%nat end). Section laws. Global Instance: SgOp (Term0 sig nat tt) := λ x, App sig _ _ _ (App sig _ _ _ (Op sig nat mult) x). Local Notation x := (Var sig nat 0%nat tt). Local Notation y := (Var sig nat 1%nat tt). Local Notation z := (Var sig nat 2%nat tt). Import notations. Inductive Laws: EqEntailment sig → Prop := | e_mult_assoc: Laws (x & (y & z) === (x & y) & z). End laws. Definition theory: EquationalTheory := Build_EquationalTheory sig Laws. Definition Object := varieties.Object theory. Definition forget: Object → setoids.Object := @product.project unit (λ _, setoids.Object) (λ _, _: Arrows setoids.Object) _ (λ _, _: CatId setoids.Object) (λ _, _: CatComp setoids.Object) (λ _, _: Category setoids.Object) tt ∘ forget_algebra.object theory ∘ forget_variety.forget theory. (* todo: too ugly *) (* Now follow a series of encoding/decoding functions to convert between the specialized Monoid/Monoid_Morphism type classes and the universal Algebra/InVariety/HomoMorphism type classes instantiated with the above signature and theory. *) #[global] Instance encode_operations A `{!SgOp A}: AlgebraOps sig (λ _, A) := λ o, match o with mult => (&) end. Section decode_operations. Context `{AlgebraOps theory A}. Global Instance: SgOp (A tt) := algebra_op mult. End decode_operations. Section encode_variety_and_ops. Context A `{SemiGroup A}. Global Instance encode_algebra_and_ops: Algebra sig _. Proof. constructor; intros []; simpl; apply _. Qed. Global Instance encode_variety_and_ops: InVariety theory (λ _, A). Proof. constructor. apply _. intros ? [] ?; simpl; unfold algebra_op; simpl. apply associativity. Qed. Definition object: Object := varieties.object theory (λ _, A). End encode_variety_and_ops. Lemma encode_algebra_only `{!AlgebraOps theory A} `{∀ u, Equiv (A u)} `{!SemiGroup (A tt)}: Algebra theory A . Proof. constructor; intros []; simpl in *; try apply _. Qed. Global Instance decode_variety_and_ops `{InVariety theory A}: SemiGroup (A tt). Proof with simpl; auto. pose proof (λ law lawgood x y z, variety_laws law lawgood (λ s n, match s with tt => match n with 0 => x | 1 => y | _ => z end end)) as laws. constructor; try apply _. intro. apply_simplified (laws _ e_mult_assoc). apply_simplified (algebra_propers mult)... Qed. Lemma encode_morphism_only `{AlgebraOps theory A} `{∀ u, Equiv (A u)} `{AlgebraOps theory B} `{∀ u, Equiv (B u)} (f: ∀ u, A u → B u) `{!SemiGroup_Morphism (f tt)}: HomoMorphism sig A B f. Proof. pose proof (sgmor_a (f:=f tt)). pose proof (sgmor_b (f:=f tt)). constructor. intros []. apply _. intros []; simpl. apply preserves_sg_op. apply (encode_algebra_only). apply encode_algebra_only. Qed. Lemma encode_morphism_and_ops `{SemiGroup_Morphism A B f}: @HomoMorphism sig (λ _, A) (λ _, B) _ _ ( _) ( _) (λ _, f). Proof. intros. apply encode_morphism_only; assumption. Qed. Lemma decode_morphism_and_ops `{InVariety theory x} `{InVariety theory y} `{!HomoMorphism theory x y f}: SemiGroup_Morphism (f tt). Proof. constructor; try apply _. apply (preserves theory x y f mult). Qed. #[global] Instance id_sg_morphism `{SemiGroup A}: SemiGroup_Morphism (@id A). Proof. repeat (split; try apply _); easy. Qed. (* Finally, we use these encoding/decoding functions to specialize some universal results: *) Section specialized. Context `{Equiv A} `{SgOp A} `{Equiv B} `{SgOp B} `{Equiv C} `{SgOp C} (f : A → B) (g : B → C). Instance compose_sg_morphism: SemiGroup_Morphism f → SemiGroup_Morphism g → SemiGroup_Morphism (g ∘ f). Proof. intros. pose proof (encode_morphism_and_ops (f:=f)) as P. pose proof (encode_morphism_and_ops (f:=g)) as Q. pose proof (@compose_homomorphisms theory _ _ _ _ _ _ _ _ _ _ _ P Q) as PP. pose proof (sgmor_a (f:=f)). pose proof (sgmor_b (f:=f)). pose proof (sgmor_b (f:=g)). apply (@decode_morphism_and_ops _ _ _ _ _ _ _ _ _ PP). Qed. Instance invert_sg_morphism: ∀ `{!Inverse f}, Bijective f → SemiGroup_Morphism f → SemiGroup_Morphism (f⁻¹). Proof. intros. pose proof (encode_morphism_and_ops (f:=f)) as P. pose proof (@invert_homomorphism theory _ _ _ _ _ _ _ _ _ _ P) as Q. pose proof (sgmor_a (f:=f)). pose proof (sgmor_b (f:=f)). apply (@decode_morphism_and_ops _ _ _ _ _ _ _ _ _ Q). Qed. End specialized. #[global] Hint Extern 4 (SemiGroup_Morphism (_ ∘ _)) => class_apply @compose_sg_morphism : typeclass_instances. #[global] Hint Extern 4 (SemiGroup_Morphism (_⁻¹)) => class_apply @invert_sg_morphism : typeclass_instances. math-classes-8.19.0/varieties/semirings.v000066400000000000000000000114561460576051100203730ustar00rootroot00000000000000(* To be imported qualified. *) Require MathClasses.theory.rings MathClasses.categories.varieties. Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms MathClasses.misc.workaround_tactics. Inductive op := plus | mult | zero | one. Definition sig: Signature := single_sorted_signature (λ o, match o with zero | one => O | plus | mult => 2%nat end). Section laws. Global Instance: Plus (Term0 sig nat tt) := λ x, App sig _ _ _ (App sig _ _ _ (Op sig _ plus) x). Global Instance: Mult (Term0 sig nat tt) := λ x, App sig _ _ _ (App sig _ _ _ (Op sig _ mult) x). Global Instance: Zero (Term0 sig nat tt) := Op sig _ zero. Global Instance: One (Term0 sig nat tt) := Op sig _ one. Local Notation x := (Var sig _ 0%nat tt). Local Notation y := (Var sig _ 1%nat tt). Local Notation z := (Var sig _ 2%nat tt). Import notations. Inductive Laws: EqEntailment sig → Prop := |e_plus_assoc: Laws (x + (y + z) === (x + y) + z) |e_plus_comm: Laws (x + y === y + x) |e_plus_0_l: Laws (0 + x === x) |e_mult_assoc: Laws (x * (y * z) === (x * y) * z) |e_mult_comm: Laws (x * y === y * x) |e_mult_1_l: Laws (1 * x === x) |e_mult_0_l: Laws (0 * x === 0) |e_distr_l: Laws (x * (y + z) === x * y + x * z) |e_distr_r: Laws ((x + y) * z === x * z + y * z). End laws. Definition theory: EquationalTheory := Build_EquationalTheory sig Laws. (* Given a SemiRing, we can make the corresponding Implementation, prove the laws, and construct the categorical object: *) Section from_instance. Context A `{SemiRing A}. Instance implementation: AlgebraOps sig (λ _, A) := λ o, match o with plus => (+) | mult => (.*.) | zero => 0: A | one => 1:A end. Global Instance: Algebra sig _. Proof. constructor. intro. apply _. intro o. destruct o; simpl; try apply _; apply reflexivity. Qed. Lemma laws en (l: Laws en) vars: eval_stmt sig vars en. Proof. inversion_clear l; simpl. apply associativity. apply commutativity. apply theory.rings.plus_0_l. apply associativity. apply commutativity. apply theory.rings.mult_1_l. unfold algebra_op. simpl. apply left_absorb. apply distribute_l. apply distribute_r. Qed. Global Instance variety: InVariety theory (λ _, A). Proof. constructor. apply _. exact laws. Qed. Definition Object := varieties.Object theory. Definition object: Object := varieties.object theory (λ _, A). End from_instance. (* Similarly, given a categorical object, we can make the corresponding class instances: *) Section ops_from_alg_to_sr. Context `{AlgebraOps theory A}. Global Instance: Plus (A tt) := algebra_op plus. Global Instance: Mult (A tt) := algebra_op mult. Global Instance: Zero (A tt) := algebra_op zero. Global Instance: One (A tt) := algebra_op one. End ops_from_alg_to_sr. Lemma mor_from_sr_to_alg `{InVariety theory A} `{InVariety theory B} (f: ∀ u, A u → B u) `{!SemiRing_Morphism (f tt)}: HomoMorphism sig A B f. Proof. constructor. intros []. apply _. intros []; simpl. apply rings.preserves_plus. apply rings.preserves_mult. change (f tt 0 = 0). apply rings.preserves_0. change (f tt 1 = 1). apply rings.preserves_1. change (Algebra theory A). apply _. change (Algebra theory B). apply _. Qed. (* todo: these [change]s should not be necessary at all. [apply] is too weak. report bug. *) #[global] Instance decode_variety_and_ops `{v: InVariety theory A}: SemiRing (A tt). Proof with simpl; auto. pose proof (λ law lawgood x y z, variety_laws law lawgood (λ s n, match s with tt => match n with 0 => x | 1 => y | _ => z end end)). repeat (constructor; try apply _); repeat intro. apply_simplified (H _ e_plus_assoc). apply_simplified (algebra_propers plus)... apply_simplified (H _ e_plus_0_l)... transitivity (algebra_op plus (algebra_op zero) x). apply_simplified (H _ e_plus_comm)... apply_simplified (H _ e_plus_0_l)... apply_simplified (H _ e_plus_comm)... apply_simplified (H _ e_mult_assoc)... apply_simplified (algebra_propers mult)... apply_simplified (H _ e_mult_1_l)... transitivity (algebra_op mult (algebra_op one) x). apply_simplified (H _ e_mult_comm)... apply_simplified (H _ e_mult_1_l)... apply_simplified (H _ e_mult_comm)... apply_simplified (H _ e_distr_l)... apply_simplified (H _ e_mult_0_l)... Qed. (* todo: clean up ring in the same way *) Lemma decode_morphism_and_ops `{InVariety theory x} `{InVariety theory y} `{!HomoMorphism theory x y f}: SemiRing_Morphism (f tt). Proof. pose proof (preserves theory x y f) as P. repeat (constructor; try apply _) ; [ apply (P plus) | apply (P zero) | apply (P mult) | apply (P one) ]. Qed. math-classes-8.19.0/varieties/setoids.v000066400000000000000000000026401460576051100200400ustar00rootroot00000000000000Require MathClasses.categories.varieties. Require Import MathClasses.interfaces.abstract_algebra MathClasses.interfaces.universal_algebra MathClasses.theory.ua_homomorphisms. Definition op := False. Definition sig: Signature := Build_Signature unit op (False_rect _). Definition Laws: EqEntailment sig → Prop := λ _, False. Definition theory: EquationalTheory := Build_EquationalTheory sig Laws. (* Constructing an object in the variety from an instance of the class: *) Section from_instance. Context A `{Setoid A}. Notation carriers := (λ _, A). Instance: AlgebraOps sig carriers := λ o, False_rect _ o. Instance: Algebra sig carriers. Proof. constructor; intuition. Qed. Instance: InVariety theory carriers. Proof. constructor; intuition. Qed. Definition object: varieties.Object theory := varieties.object theory (λ _, A). End from_instance. (* Constructing an instance of the class from an object in the variety: *) Section from_variety. Context `{InVariety theory A}. Instance: Setoid (A tt) := {}. End from_variety. Lemma mor_from_sr_to_alg `{InVariety theory A} `{InVariety theory B} (f: ∀ u, A u → B u) `{!Setoid_Morphism (f tt)}: HomoMorphism sig A B f. Proof with try apply _. constructor. intros []... intros []. change (Algebra theory A)... change (Algebra theory B)... Qed. #[global] Instance decode_variety_and_ops `{InVariety theory A}: Setoid (A tt) := {}.