src/HOL/GCD.thy
Mon, 18 Apr 2016 20:56:13 +0200 haftmann capitalized GCD and LCM syntax
Fri, 26 Feb 2016 22:15:09 +0100 Manuel Eberl Tuned Euclidean Rings/GCD rings
Wed, 17 Feb 2016 21:51:58 +0100 haftmann dropped various legacy fact bindings and tuned proofs
Wed, 17 Feb 2016 21:51:57 +0100 haftmann more sophisticated GCD syntax
Wed, 17 Feb 2016 21:51:57 +0100 haftmann cleansed junk-producing interpretations for gcd/lcm on nat altogether
Wed, 17 Feb 2016 21:51:57 +0100 haftmann dropped various legacy fact bindings
Wed, 17 Feb 2016 21:51:57 +0100 haftmann generalized some lemmas;
Wed, 17 Feb 2016 21:51:56 +0100 haftmann more theorems concerning gcd/lcm/Gcd/Lcm
Wed, 17 Feb 2016 21:51:56 +0100 haftmann further generalization and polishing
Wed, 17 Feb 2016 21:51:56 +0100 haftmann pulled out legacy aliasses and infamous dvd interpretations into theory appendix
Wed, 17 Feb 2016 21:51:56 +0100 haftmann prefer abbreviations for compound operators INFIMUM and SUPREMUM
Wed, 30 Dec 2015 11:37:29 +0100 wenzelm isabelle update_cartouches -c -t;
Mon, 28 Dec 2015 19:23:15 +0100 wenzelm more symbols;
Mon, 28 Dec 2015 01:26:34 +0100 wenzelm prefer symbols for "abs";
Thu, 24 Dec 2015 09:42:49 +0100 haftmann tuned proofs and augmented lemmas
Tue, 22 Dec 2015 15:38:59 +0100 haftmann tuned proofs and augmented some lemmas
Fri, 18 Dec 2015 11:14:28 +0100 Andreas Lochbihler add gcd instance for integer and serialisation to target language operations
Mon, 07 Dec 2015 10:38:04 +0100 wenzelm isabelle update_cartouches -c -t;
Fri, 13 Nov 2015 12:27:13 +0000 paulson Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
Mon, 09 Nov 2015 15:48:17 +0100 wenzelm qualifier is mandatory by default;
Wed, 04 Nov 2015 08:13:52 +0100 ballarin Keyword 'rewrites' identifies rewrite morphisms.
Sun, 13 Sep 2015 22:56:52 +0200 wenzelm tuned proofs -- less legacy;
Sat, 18 Jul 2015 22:58:50 +0200 wenzelm isabelle update_cartouches;
Wed, 08 Jul 2015 20:19:12 +0200 haftmann tuned facts
Wed, 08 Jul 2015 14:01:41 +0200 haftmann more cautious use of [iff] declarations
Wed, 08 Jul 2015 14:01:41 +0200 haftmann avoid explicit definition of the relation of associated elements in a ring -- prefer explicit normalization instead
Wed, 08 Jul 2015 14:01:40 +0200 haftmann eliminated some duplication
Wed, 08 Jul 2015 14:01:39 +0200 haftmann more algebraic properties for gcd/lcm
Sat, 27 Jun 2015 20:20:33 +0200 haftmann tuned code setup
Sat, 27 Jun 2015 20:20:32 +0200 haftmann algebraic specification for set gcd
Thu, 25 Jun 2015 23:33:47 +0200 wenzelm tuned proofs;
Wed, 17 Jun 2015 23:01:19 +0200 wenzelm tuned proofs -- slightly faster;
Tue, 02 Jun 2015 09:10:05 +0200 wenzelm tuned proof;
Thu, 30 Apr 2015 15:28:01 +0100 paulson tidying some messy proofs
Wed, 08 Apr 2015 21:48:59 +0200 wenzelm eliminated hard tabs;
Wed, 25 Mar 2015 10:44:57 +0100 wenzelm prefer local fixes;
Tue, 10 Mar 2015 15:20:40 +0000 paulson Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
Sun, 15 Feb 2015 17:01:22 +0100 haftmann explicit equivalence for strict order on lattices
Tue, 10 Feb 2015 14:29:36 +0100 wenzelm indicate slow proof (approx. 20s);
Mon, 17 Nov 2014 14:55:32 +0100 haftmann formally self-contained gcd type classes
Sun, 02 Nov 2014 18:21:45 +0100 wenzelm modernized header uniformly as section;
Thu, 30 Oct 2014 21:02:01 +0100 haftmann more simp rules concerning dvd and even/odd
Sun, 26 Oct 2014 19:11:16 +0100 haftmann eliminated redundancies;
Fri, 24 Oct 2014 15:07:51 +0200 hoelzl use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
Thu, 23 Oct 2014 14:04:05 +0200 haftmann downshift of theory Parity in the hierarchy
Tue, 07 Oct 2014 23:29:43 +0200 wenzelm more bibtex entries;
Sat, 05 Jul 2014 11:01:53 +0200 haftmann prefer ac_simps collections over separate name bindings for add and mult
Fri, 04 Jul 2014 20:18:47 +0200 haftmann reduced name variants for assoc and commute on plus and mult
Wed, 19 Mar 2014 18:47:22 +0100 haftmann elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
Sun, 16 Mar 2014 18:09:04 +0100 haftmann normalising simp rules for compound operators
Thu, 26 Dec 2013 22:47:49 +0100 haftmann prefer ephemeral interpretation over interpretation in proof contexts;
Tue, 19 Nov 2013 10:05:53 +0100 haftmann eliminiated neg_numeral in favour of - (numeral _)
Fri, 15 Nov 2013 22:02:01 +0100 haftmann proper code equations for Gcd and Lcm on nat and int
Tue, 05 Nov 2013 09:44:57 +0100 hoelzl generalize SUP and INF to the syntactic type classes Sup and Inf
Thu, 25 Jul 2013 08:57:16 +0200 haftmann factored syntactic type classes for bot and top (by Alessandro Coglio)
Wed, 19 Jun 2013 17:33:51 +0200 noschinl added coprimality lemma
Tue, 26 Mar 2013 22:09:39 +0100 haftmann avoid odd foundational terms after interpretation;
Sat, 23 Mar 2013 20:50:39 +0100 haftmann fundamental revision of big operators on sets
Fri, 19 Oct 2012 15:12:52 +0200 webertj Renamed {left,right}_distrib to distrib_{right,left}.
Fri, 27 Jul 2012 19:57:23 +0200 wenzelm tuned proofs -- avoid odd situations of polymorphic Frees in goal state;
less more (0) -100 -60 tip