src/HOL/Library/Parallel.thy
author wenzelm
Sun, 27 May 2018 22:37:08 +0200
changeset 68300 cd8ab1a7a286
parent 60500 903bb1495239
child 81706 7beb0cf38292
permissions -rw-r--r--
retain isolated blob nodes (amending deb2fcbda16e): avoid failure of Session.handle_change with "Missing blob", when opening theory with load command later;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48427
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
     2
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
     3
section \<open>Futures and parallel lists 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
     4
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
     5
theory Parallel
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
     6
imports Main
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
     7
begin
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
     8
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
     9
subsection \<open>Futures\<close>
48427
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    10
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    11
datatype 'a future = fork "unit \<Rightarrow> 'a"
48427
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    12
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    13
primrec join :: "'a future \<Rightarrow> 'a" where
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    14
  "join (fork f) = f ()"
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    15
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    16
lemma future_eqI [intro!]:
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    17
  assumes "join f = join g"
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    18
  shows "f = g"
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    19
  using assms by (cases f, cases g) (simp add: ext)
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    20
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 48427
diff changeset
    21
code_printing
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 48427
diff changeset
    22
  type_constructor future \<rightharpoonup> (Eval) "_ future"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 48427
diff changeset
    23
| constant fork \<rightharpoonup> (Eval) "Future.fork"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 48427
diff changeset
    24
| constant join \<rightharpoonup> (Eval) "Future.join"
48427
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    25
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    26
code_reserved Eval Future future
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    27
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    28
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    29
subsection \<open>Parallel lists\<close>
48427
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
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    32
  [simp]: "map = List.map"
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    33
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    34
definition forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    35
  "forall = list_all"
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    36
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    37
lemma forall_all [simp]:
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    38
  "forall P xs \<longleftrightarrow> (\<forall>x\<in>set xs. P x)"
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    39
  by (simp add: forall_def list_all_iff)
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    40
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    41
definition exists :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    42
  "exists = list_ex"
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    43
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    44
lemma exists_ex [simp]:
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    45
  "exists P xs \<longleftrightarrow> (\<exists>x\<in>set xs. P x)"
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    46
  by (simp add: exists_def list_ex_iff)
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    47
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 48427
diff changeset
    48
code_printing
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 48427
diff changeset
    49
  constant map \<rightharpoonup> (Eval) "Par'_List.map"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 48427
diff changeset
    50
| constant forall \<rightharpoonup> (Eval) "Par'_List.forall"
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 48427
diff changeset
    51
| constant exists \<rightharpoonup> (Eval) "Par'_List.exists"
48427
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    52
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    53
code_reserved Eval Par_List
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
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    56
hide_const (open) fork join map exists forall
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
end
571cb1df0768 library theories for debugging and parallel computing using code generation towards Isabelle/ML
haftmann
parents:
diff changeset
    59