| author | blanchet | 
| Fri, 15 Feb 2013 09:17:26 +0100 | |
| changeset 51138 | 583a37b9512d | 
| parent 45169 | 4baa475a51e6 | 
| child 51272 | 9c8d63b4b6be | 
| permissions | -rw-r--r-- | 
| 45047 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 1 | (* Title: HOL/Proofs/Extraction/Higman_Extraction.thy | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 2 | Author: Stefan Berghofer, TU Muenchen | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 3 | Author: Monika Seisenberger, LMU Muenchen | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 4 | *) | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 5 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 6 | (*<*) | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 7 | theory Higman_Extraction | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 8 | imports Higman "~~/src/HOL/Library/State_Monad" Random | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 9 | begin | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 10 | (*>*) | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 11 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 12 | subsection {* Extracting the program *}
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 13 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 14 | declare R.induct [ind_realizer] | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 15 | declare T.induct [ind_realizer] | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 16 | declare L.induct [ind_realizer] | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 17 | declare good.induct [ind_realizer] | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 18 | declare bar.induct [ind_realizer] | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 19 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 20 | extract higman_idx | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 21 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 22 | text {*
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 23 |   Program extracted from the proof of @{text higman_idx}:
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 24 |   @{thm [display] higman_idx_def [no_vars]}
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 25 | Corresponding correctness theorem: | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 26 |   @{thm [display] higman_idx_correctness [no_vars]}
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 27 |   Program extracted from the proof of @{text higman}:
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 28 |   @{thm [display] higman_def [no_vars]}
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 29 |   Program extracted from the proof of @{text prop1}:
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 30 |   @{thm [display] prop1_def [no_vars]}
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 31 |   Program extracted from the proof of @{text prop2}:
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 32 |   @{thm [display] prop2_def [no_vars]}
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 33 |   Program extracted from the proof of @{text prop3}:
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 34 |   @{thm [display] prop3_def [no_vars]}
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 35 | *} | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 36 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 37 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 38 | subsection {* Some examples *}
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 39 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 40 | instantiation LT and TT :: default | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 41 | begin | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 42 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 43 | definition "default = L0 [] []" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 44 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 45 | definition "default = T0 A [] [] [] R0" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 46 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 47 | instance .. | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 48 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 49 | end | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 50 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 51 | function mk_word_aux :: "nat \<Rightarrow> Random.seed \<Rightarrow> letter list \<times> Random.seed" where | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 52 |   "mk_word_aux k = exec {
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 53 | i \<leftarrow> Random.range 10; | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 54 | (if i > 7 \<and> k > 2 \<or> k > 1000 then Pair [] | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 55 |      else exec {
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 56 | let l = (if i mod 2 = 0 then A else B); | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 57 | ls \<leftarrow> mk_word_aux (Suc k); | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 58 | Pair (l # ls) | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 59 | })}" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 60 | by pat_completeness auto | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 61 | termination by (relation "measure ((op -) 1001)") auto | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 62 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 63 | definition mk_word :: "Random.seed \<Rightarrow> letter list \<times> Random.seed" where | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 64 | "mk_word = mk_word_aux 0" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 65 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 66 | primrec mk_word_s :: "nat \<Rightarrow> Random.seed \<Rightarrow> letter list \<times> Random.seed" where | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 67 | "mk_word_s 0 = mk_word" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 68 |   | "mk_word_s (Suc n) = exec {
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 69 | _ \<leftarrow> mk_word; | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 70 | mk_word_s n | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 71 | }" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 72 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 73 | definition g1 :: "nat \<Rightarrow> letter list" where | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 74 | "g1 s = fst (mk_word_s s (20000, 1))" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 75 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 76 | definition g2 :: "nat \<Rightarrow> letter list" where | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 77 | "g2 s = fst (mk_word_s s (50000, 1))" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 78 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 79 | fun f1 :: "nat \<Rightarrow> letter list" where | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 80 | "f1 0 = [A, A]" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 81 | | "f1 (Suc 0) = [B]" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 82 | | "f1 (Suc (Suc 0)) = [A, B]" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 83 | | "f1 _ = []" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 84 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 85 | fun f2 :: "nat \<Rightarrow> letter list" where | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 86 | "f2 0 = [A, A]" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 87 | | "f2 (Suc 0) = [B]" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 88 | | "f2 (Suc (Suc 0)) = [B, A]" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 89 | | "f2 _ = []" | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 90 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 91 | ML {*
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 92 | local | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 93 |   val higman_idx = @{code higman_idx};
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 94 |   val g1 = @{code g1};
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 95 |   val g2 = @{code g2};
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 96 |   val f1 = @{code f1};
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 97 |   val f2 = @{code f2};
 | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 98 | in | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 99 | val (i1, j1) = higman_idx g1; | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 100 | val (v1, w1) = (g1 i1, g1 j1); | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 101 | val (i2, j2) = higman_idx g2; | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 102 | val (v2, w2) = (g2 i2, g2 j2); | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 103 | val (i3, j3) = higman_idx f1; | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 104 | val (v3, w3) = (f1 i3, f1 j3); | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 105 | val (i4, j4) = higman_idx f2; | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 106 | val (v4, w4) = (f2 i4, f2 j4); | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 107 | end; | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 108 | *} | 
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 109 | |
| 
3aa8d3c391a4
Moved extraction part of Higman's lemma to separate theory to allow reuse in
 berghofe parents: diff
changeset | 110 | end |