merged
authorwenzelm
Sun, 28 Feb 2016 21:58:06 +0100
changeset 62461 075ef5ec115c
parent 62458 9590972c2caf (diff)
parent 62460 4b2018eb92e8 (current diff)
child 62462 c7def2433a06
child 62464 08e62096e7f4
merged
lib/scripts/run-polyml
lib/scripts/run-polyml-5.5.1
lib/scripts/run-polyml-5.5.2
lib/scripts/run-polyml-5.6
src/Pure/ML/ml_statistics_polyml-5.5.0.ML
src/Pure/RAW/ROOT_polyml-5.5.2.ML
src/Pure/RAW/exn_trace_polyml-5.5.1.ML
--- a/src/HOL/Number_Theory/Euclidean_Algorithm.thy	Sun Feb 28 21:25:55 2016 +0100
+++ b/src/HOL/Number_Theory/Euclidean_Algorithm.thy	Sun Feb 28 21:58:06 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))"