author | nipkow |
Thu, 31 Aug 2017 09:50:11 +0200 | |
changeset 66566 | a14bbbaa628d |
parent 66453 | cc19f7ca2ed6 |
child 68484 | 59793df7f853 |
permissions | -rw-r--r-- |
31378 | 1 |
|
2 |
(* Author: Florian Haftmann, TU Muenchen *) |
|
21917 | 3 |
|
63167 | 4 |
section \<open>A huge collection of equations to generate code from\<close> |
21917 | 5 |
|
37695
71e84a203c19
introduced distinct session HOL-Codegenerator_Test
haftmann
parents:
37407
diff
changeset
|
6 |
theory Candidates |
21917 | 7 |
imports |
27421 | 8 |
Complex_Main |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66251
diff
changeset
|
9 |
"HOL-Library.Library" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66251
diff
changeset
|
10 |
"HOL-Library.Subseq_Order" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66251
diff
changeset
|
11 |
"HOL-Library.RBT" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66251
diff
changeset
|
12 |
"HOL-Data_Structures.Tree_Map" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66251
diff
changeset
|
13 |
"HOL-Data_Structures.Tree_Set" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66251
diff
changeset
|
14 |
"HOL-Computational_Algebra.Computational_Algebra" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66251
diff
changeset
|
15 |
"HOL-Computational_Algebra.Polynomial_Factorial" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66251
diff
changeset
|
16 |
"HOL-Number_Theory.Eratosthenes" |
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66251
diff
changeset
|
17 |
"HOL-ex.Records" |
21917 | 18 |
begin |
19 |
||
61129
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
20 |
text \<open>Drop technical stuff from @{theory Quickcheck_Narrowing} which is tailored towards Haskell\<close> |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
21 |
|
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
58889
diff
changeset
|
22 |
setup \<open> |
59842 | 23 |
fn thy => |
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
58889
diff
changeset
|
24 |
let |
59842 | 25 |
val tycos = Sign.logical_types thy; |
26 |
val consts = map_filter (try (curry (Axclass.param_of_inst thy) |
|
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
58889
diff
changeset
|
27 |
@{const_name "Quickcheck_Narrowing.partial_term_of"})) tycos; |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66014
diff
changeset
|
28 |
in fold Code.declare_unimplemented_global consts thy end |
61129
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
29 |
\<close> |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
30 |
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
31 |
text \<open>Simple example for the predicate compiler.\<close> |
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
58889
diff
changeset
|
32 |
|
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51160
diff
changeset
|
33 |
inductive sublist :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" |
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51160
diff
changeset
|
34 |
where |
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51160
diff
changeset
|
35 |
empty: "sublist [] xs" |
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51160
diff
changeset
|
36 |
| drop: "sublist ys xs \<Longrightarrow> sublist ys (x # xs)" |
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51160
diff
changeset
|
37 |
| take: "sublist ys xs \<Longrightarrow> sublist (x # ys) (x # xs)" |
33500 | 38 |
|
39 |
code_pred sublist . |
|
40 |
||
61129
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
41 |
text \<open>Avoid popular infix.\<close> |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
42 |
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
43 |
code_reserved SML upto |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
44 |
|
63167 | 45 |
text \<open>Explicit check in \<open>OCaml\<close> for correct precedence of let expressions in list expressions\<close> |
61129
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
46 |
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
47 |
definition funny_list :: "bool list" |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
48 |
where |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
49 |
"funny_list = [let b = True in b, False]" |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
50 |
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
51 |
definition funny_list' :: "bool list" |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
52 |
where |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
53 |
"funny_list' = funny_list" |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
54 |
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
55 |
lemma [code]: |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
56 |
"funny_list' = [True, False]" |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
57 |
by (simp add: funny_list_def funny_list'_def) |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
58 |
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
59 |
definition check_list :: unit |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
60 |
where |
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
61 |
"check_list = (if funny_list = funny_list' then () else undefined)" |
31378 | 62 |
|
63167 | 63 |
text \<open>Explicit check in \<open>Scala\<close> for correct bracketing of abstractions\<close> |
61130
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
61129
diff
changeset
|
64 |
|
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
61129
diff
changeset
|
65 |
definition funny_funs :: "(bool \<Rightarrow> bool) list \<Rightarrow> (bool \<Rightarrow> bool) list" |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
61129
diff
changeset
|
66 |
where |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
61129
diff
changeset
|
67 |
"funny_funs fs = (\<lambda>x. x \<or> True) # (\<lambda>x. x \<or> False) # fs" |
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
61129
diff
changeset
|
68 |
|
21917 | 69 |
end |