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