| 13405 |      1 | (*  Title:      HOL/Extraction/QuotRem.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Stefan Berghofer, TU Muenchen
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | header {* Quotient and remainder *}
 | 
|  |      7 | 
 | 
|  |      8 | theory QuotRem = Main:
 | 
|  |      9 | 
 | 
|  |     10 | text {* Derivation of quotient and remainder using program extraction. *}
 | 
|  |     11 | 
 | 
| 13931 |     12 | lemma nat_eq_dec: "\<And>n::nat. m = n \<or> m \<noteq> n"
 | 
|  |     13 |   apply (induct m)
 | 
|  |     14 |   apply (case_tac n)
 | 
|  |     15 |   apply (case_tac [3] na)
 | 
|  |     16 |   apply (simp only: nat.simps, rules?)+
 | 
| 13405 |     17 |   done
 | 
|  |     18 | 
 | 
| 13931 |     19 | theorem division: "\<exists>r q. a = Suc b * q + r \<and> r \<le> b"
 | 
|  |     20 | proof (induct a)
 | 
|  |     21 |   case 0
 | 
|  |     22 |   have "0 = Suc b * 0 + 0 \<and> 0 \<le> b" by simp
 | 
|  |     23 |   thus ?case by rules
 | 
|  |     24 | next
 | 
|  |     25 |   case (Suc a)
 | 
|  |     26 |   then obtain r q where I: "a = Suc b * q + r" and "r \<le> b" by rules
 | 
|  |     27 |   from nat_eq_dec show ?case
 | 
|  |     28 |   proof
 | 
|  |     29 |     assume "r = b"
 | 
|  |     30 |     with I have "Suc a = Suc b * (Suc q) + 0 \<and> 0 \<le> b" by simp
 | 
|  |     31 |     thus ?case by rules
 | 
|  |     32 |   next
 | 
|  |     33 |     assume "r \<noteq> b"
 | 
|  |     34 |     hence "r < b" by (simp add: order_less_le)
 | 
|  |     35 |     with I have "Suc a = Suc b * q + (Suc r) \<and> (Suc r) \<le> b" by simp
 | 
|  |     36 |     thus ?case by rules
 | 
|  |     37 |   qed
 | 
|  |     38 | qed
 | 
| 13405 |     39 | 
 | 
|  |     40 | extract division
 | 
|  |     41 | 
 | 
|  |     42 | text {*
 | 
|  |     43 |   The program extracted from the above proof looks as follows
 | 
|  |     44 |   @{thm [display] division_def [no_vars]}
 | 
|  |     45 |   The corresponding correctness theorem is
 | 
|  |     46 |   @{thm [display] division_correctness [no_vars]}
 | 
|  |     47 | *}
 | 
|  |     48 | 
 | 
|  |     49 | generate_code
 | 
|  |     50 |   test = "division 9 2"
 | 
|  |     51 | 
 | 
|  |     52 | end
 |