| 45610 |      1 | (*  Title:      HOL/SPARK/Manual/Simple_Greatest_Common_Divisor.thy
 | 
| 45044 |      2 |     Author:     Stefan Berghofer
 | 
|  |      3 |     Copyright:  secunet Security Networks AG
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | theory Simple_Greatest_Common_Divisor
 | 
|  |      7 | imports SPARK GCD
 | 
|  |      8 | begin
 | 
|  |      9 | 
 | 
|  |     10 | spark_proof_functions
 | 
|  |     11 |   gcd = "gcd :: int \<Rightarrow> int \<Rightarrow> int"
 | 
|  |     12 | 
 | 
|  |     13 | spark_open "simple_greatest_common_divisor/g_c_d.siv"
 | 
|  |     14 | 
 | 
|  |     15 | spark_vc procedure_g_c_d_4
 | 
|  |     16 |   using `0 < d` `gcd c d = gcd m n`
 | 
|  |     17 |   by (simp add: gcd_non_0_int)
 | 
|  |     18 | 
 | 
|  |     19 | spark_vc procedure_g_c_d_9
 | 
|  |     20 |   using `0 \<le> c` `gcd c 0 = gcd m n`
 | 
|  |     21 |   by simp
 | 
|  |     22 | 
 | 
|  |     23 | spark_end
 | 
|  |     24 | 
 | 
|  |     25 | end
 |