| author | wenzelm | 
| Mon, 29 May 2017 19:34:07 +0200 | |
| changeset 65959 | 47309113ee4d | 
| parent 63882 | 018998c00003 | 
| child 66453 | cc19f7ca2ed6 | 
| permissions | -rw-r--r-- | 
| 61343 | 1 | section \<open>A simple example demonstrating parallelism for code generated towards Isabelle/ML\<close> | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 2 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 3 | theory Parallel_Example | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 4 | imports Complex_Main "~~/src/HOL/Library/Parallel" "~~/src/HOL/Library/Debug" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 5 | begin | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 6 | |
| 61343 | 7 | subsection \<open>Compute-intensive examples.\<close> | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 8 | |
| 61343 | 9 | subsubsection \<open>Fragments of the harmonic series\<close> | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 10 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 11 | definition harmonic :: "nat \<Rightarrow> rat" where | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
61933diff
changeset | 12 | "harmonic n = sum_list (map (\<lambda>n. 1 / of_nat n) [1..<n])" | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 13 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 14 | |
| 61343 | 15 | subsubsection \<open>The sieve of Erathostenes\<close> | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 16 | |
| 61343 | 17 | text \<open> | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 18 | The attentive reader may relate this ad-hoc implementation to the | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 19 | arithmetic notion of prime numbers as a little exercise. | 
| 61343 | 20 | \<close> | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 21 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 22 | primrec mark :: "nat \<Rightarrow> nat \<Rightarrow> bool list \<Rightarrow> bool list" where | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 23 | "mark _ _ [] = []" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 24 | | "mark m n (p # ps) = (case n of 0 \<Rightarrow> False # mark m m ps | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 25 | | Suc n \<Rightarrow> p # mark m n ps)" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 26 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 27 | lemma length_mark [simp]: | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 28 | "length (mark m n ps) = length ps" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 29 | by (induct ps arbitrary: n) (simp_all split: nat.split) | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 30 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 31 | function sieve :: "nat \<Rightarrow> bool list \<Rightarrow> bool list" where | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 32 | "sieve m ps = (case dropWhile Not ps | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 33 | of [] \<Rightarrow> ps | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 34 | | p#ps' \<Rightarrow> let n = m - length ps' in takeWhile Not ps @ p # sieve m (mark n n ps'))" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 35 | by pat_completeness auto | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 36 | |
| 61933 | 37 | termination \<comment> \<open>tuning of this proof is left as an exercise to the reader\<close> | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 38 | apply (relation "measure (length \<circ> snd)") | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 39 | apply rule | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 40 | apply (auto simp add: length_dropWhile_le) | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 41 | proof - | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 42 | fix ps qs q | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 43 | assume "dropWhile Not ps = q # qs" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 44 | then have "length (q # qs) = length (dropWhile Not ps)" by simp | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 45 | then have "length qs < length (dropWhile Not ps)" by simp | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 46 | moreover have "length (dropWhile Not ps) \<le> length ps" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 47 | by (simp add: length_dropWhile_le) | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 48 | ultimately show "length qs < length ps" by auto | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 49 | qed | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 50 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 51 | primrec natify :: "nat \<Rightarrow> bool list \<Rightarrow> nat list" where | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 52 | "natify _ [] = []" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 53 | | "natify n (p#ps) = (if p then n # natify (Suc n) ps else natify (Suc n) ps)" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 54 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 55 | primrec list_primes where | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 56 | "list_primes (Suc n) = natify 1 (sieve n (False # replicate n True))" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 57 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 58 | |
| 61343 | 59 | subsubsection \<open>Naive factorisation\<close> | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 60 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 61 | function factorise_from :: "nat \<Rightarrow> nat \<Rightarrow> nat list" where | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 62 | "factorise_from k n = (if 1 < k \<and> k \<le> n | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 63 | then | 
| 61433 
a4c0de1df3d8
qualify some names stemming from internal bootstrap constructions
 haftmann parents: 
61343diff
changeset | 64 | let (q, r) = Divides.divmod_nat n k | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 65 | in if r = 0 then k # factorise_from k q | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 66 | else factorise_from (Suc k) n | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 67 | else [])" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 68 | by pat_completeness auto | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 69 | |
| 61933 | 70 | termination factorise_from \<comment> \<open>tuning of this proof is left as an exercise to the reader\<close> | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 71 | term measure | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 72 | apply (relation "measure (\<lambda>(k, n). 2 * n - k)") | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 73 | apply (auto simp add: prod_eq_iff) | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 74 | apply (case_tac "k \<le> 2 * q") | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 75 | apply (rule diff_less_mono) | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 76 | apply auto | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 77 | done | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 78 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 79 | definition factorise :: "nat \<Rightarrow> nat list" where | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 80 | "factorise n = factorise_from 2 n" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 81 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 82 | |
| 61343 | 83 | subsection \<open>Concurrent computation via futures\<close> | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 84 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 85 | definition computation_harmonic :: "unit \<Rightarrow> rat" where | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 86 | "computation_harmonic _ = Debug.timing (STR ''harmonic example'') harmonic 300" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 87 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 88 | definition computation_primes :: "unit \<Rightarrow> nat list" where | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 89 | "computation_primes _ = Debug.timing (STR ''primes example'') list_primes 4000" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 90 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 91 | definition computation_future :: "unit \<Rightarrow> nat list \<times> rat" where | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 92 | "computation_future = Debug.timing (STR ''overall computation'') | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 93 | (\<lambda>() \<Rightarrow> let c = Parallel.fork computation_harmonic | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 94 | in (computation_primes (), Parallel.join c))" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 95 | |
| 56927 | 96 | value "computation_future ()" | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 97 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 98 | definition computation_factorise :: "nat \<Rightarrow> nat list" where | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 99 | "computation_factorise = Debug.timing (STR ''factorise'') factorise" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 100 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 101 | definition computation_parallel :: "unit \<Rightarrow> nat list list" where | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 102 | "computation_parallel _ = Debug.timing (STR ''overall computation'') | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 103 | (Parallel.map computation_factorise) [20000..<20100]" | 
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 104 | |
| 56927 | 105 | value "computation_parallel ()" | 
| 48427 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 106 | |
| 
571cb1df0768
library theories for debugging and parallel computing using code generation towards Isabelle/ML
 haftmann parents: diff
changeset | 107 | end |