author | Manuel Eberl <eberlm@in.tum.de> |
Sun, 28 Feb 2016 21:20:11 +0100 | |
changeset 62458 | 9590972c2caf |
parent 62457 | a3c7bd201da7 (diff) |
parent 62456 | 11e06f5283bc (current diff) |
child 62461 | 075ef5ec115c |
Admin/isatest/settings/at-poly-e | file | annotate | diff | comparison | revisions | |
Admin/isatest/settings/at-poly-test | file | annotate | diff | comparison | revisions | |
Admin/isatest/settings/mac-poly-M2-alternative | file | annotate | diff | comparison | revisions | |
lib/Tools/update_semicolons | file | annotate | diff | comparison | revisions | |
lib/Tools/update_sub_sup | file | annotate | diff | comparison | revisions | |
lib/Tools/yxml | file | annotate | diff | comparison | revisions | |
lib/scripts/update_sub_sup | file | annotate | diff | comparison | revisions | |
lib/scripts/yxml | file | annotate | diff | comparison | revisions | |
src/Pure/Tools/check_source.scala | file | annotate | diff | comparison | revisions | |
src/Pure/Tools/update_semicolons.scala | file | annotate | diff | comparison | revisions |
--- a/src/HOL/Number_Theory/Euclidean_Algorithm.thy Sun Feb 28 19:56:57 2016 +0100 +++ b/src/HOL/Number_Theory/Euclidean_Algorithm.thy Sun Feb 28 21:20:11 2016 +0100 @@ -234,6 +234,8 @@ class euclidean_ring = euclidean_semiring + idom begin +subclass ring_div .. + function euclid_ext_aux :: "'a \<Rightarrow> _" where "euclid_ext_aux r' r s' s t' t = ( if r = 0 then let c = 1 div unit_factor r' in (s' * c, t' * c, normalize r') @@ -305,6 +307,10 @@ "case euclid_ext x y of (a,b,c) \<Rightarrow> a * x + b * y = c \<and> c = gcd_eucl x y" unfolding euclid_ext_def by (rule euclid_ext_aux_correct) simp_all +lemma euclid_ext_gcd_eucl: + "(case euclid_ext x y of (a,b,c) \<Rightarrow> c) = gcd_eucl x y" + using euclid_ext_correct'[of x y] by (simp add: case_prod_unfold) + definition euclid_ext' where "euclid_ext' x y = (case euclid_ext x y of (a, b, _) \<Rightarrow> (a, b))"