src/HOL/Extraction/QuotRem.thy
author blanchet
Wed, 04 Mar 2009 11:05:29 +0100
changeset 30242 aea5d7fa7ef5
parent 27982 2aaa4a5569a6
child 36862 952b2b102a0a
permissions -rw-r--r--
Merge.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Extraction/QuotRem.thy
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     4
*)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     5
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     6
header {* Quotient and remainder *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     7
27436
9581777503e9 code antiquotation roaring ahead
haftmann
parents: 25419
diff changeset
     8
theory QuotRem
9581777503e9 code antiquotation roaring ahead
haftmann
parents: 25419
diff changeset
     9
imports Util
9581777503e9 code antiquotation roaring ahead
haftmann
parents: 25419
diff changeset
    10
begin
9581777503e9 code antiquotation roaring ahead
haftmann
parents: 25419
diff changeset
    11
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    12
text {* Derivation of quotient and remainder using program extraction. *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    13
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    14
theorem division: "\<exists>r q. a = Suc b * q + r \<and> r \<le> b"
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    15
proof (induct a)
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    16
  case 0
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    17
  have "0 = Suc b * 0 + 0 \<and> 0 \<le> b" by simp
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
    18
  thus ?case by iprover
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    19
next
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    20
  case (Suc a)
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
    21
  then obtain r q where I: "a = Suc b * q + r" and "r \<le> b" by iprover
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    22
  from nat_eq_dec show ?case
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    23
  proof
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    24
    assume "r = b"
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    25
    with I have "Suc a = Suc b * (Suc q) + 0 \<and> 0 \<le> b" by simp
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
    26
    thus ?case by iprover
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    27
  next
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    28
    assume "r \<noteq> b"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 22845
diff changeset
    29
    with `r \<le> b` have "r < b" by (simp add: order_less_le)
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    30
    with I have "Suc a = Suc b * q + (Suc r) \<and> (Suc r) \<le> b" by simp
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 17145
diff changeset
    31
    thus ?case by iprover
13931
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    32
  qed
38d43d168e87 Converted main proof to Isar.
berghofe
parents: 13405
diff changeset
    33
qed
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    34
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    35
extract division
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    36
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    37
text {*
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    38
  The program extracted from the above proof looks as follows
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    39
  @{thm [display] division_def [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    40
  The corresponding correctness theorem is
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    41
  @{thm [display] division_correctness [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    42
*}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    43
27982
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 27436
diff changeset
    44
lemma "division 9 2 = (0, 3)" by evaluation
27436
9581777503e9 code antiquotation roaring ahead
haftmann
parents: 25419
diff changeset
    45
lemma "division 9 2 = (0, 3)" by eval
20593
5af400cc64d5 added some stuff for code generation 2
haftmann
parents: 17604
diff changeset
    46
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    47
end