| author | nipkow |
| Tue, 19 Jan 2016 11:46:54 +0100 | |
| changeset 62204 | 7f5579b12b0a |
| parent 61227 | 19ee25fe9737 |
| child 62352 | 35a9e1cbb5b3 |
| permissions | -rw-r--r-- |
| 31378 | 1 |
|
2 |
(* Author: Florian Haftmann, TU Muenchen *) |
|
| 21917 | 3 |
|
| 58889 | 4 |
section {* A huge collection of equations to generate code from *}
|
| 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 |
|
51160
599ff65b85e2
systematic conversions between nat and nibble/char;
haftmann
parents:
49077
diff
changeset
|
9 |
"~~/src/HOL/Library/Library" |
|
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51160
diff
changeset
|
10 |
"~~/src/HOL/Library/Sublist_Order" |
|
61227
19ee25fe9737
include some data structures into code generation
haftmann
parents:
61130
diff
changeset
|
11 |
"~~/src/HOL/Data_Structures/Tree_Map" |
|
19ee25fe9737
include some data structures into code generation
haftmann
parents:
61130
diff
changeset
|
12 |
"~~/src/HOL/Data_Structures/Tree_Set" |
| 51173 | 13 |
"~~/src/HOL/Number_Theory/Eratosthenes" |
| 35303 | 14 |
"~~/src/HOL/ex/Records" |
| 21917 | 15 |
begin |
16 |
||
|
61129
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
17 |
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
|
18 |
|
|
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
58889
diff
changeset
|
19 |
setup \<open> |
| 59842 | 20 |
fn thy => |
|
59484
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents:
58889
diff
changeset
|
21 |
let |
| 59842 | 22 |
val tycos = Sign.logical_types thy; |
23 |
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
|
24 |
@{const_name "Quickcheck_Narrowing.partial_term_of"})) tycos;
|
| 59842 | 25 |
in fold Code.del_eqns 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
|
26 |
\<close> |
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
27 |
|
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
28 |
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
|
29 |
|
|
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51160
diff
changeset
|
30 |
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
|
31 |
where |
|
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51160
diff
changeset
|
32 |
empty: "sublist [] xs" |
|
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
51160
diff
changeset
|
33 |
| 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
|
34 |
| take: "sublist ys xs \<Longrightarrow> sublist (x # ys) (x # xs)" |
| 33500 | 35 |
|
36 |
code_pred sublist . |
|
37 |
||
|
61129
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
38 |
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
|
39 |
|
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
40 |
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
|
41 |
|
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
42 |
text \<open>Explicit check in @{text OCaml} for correct precedence of let expressions in list expressions\<close>
|
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
43 |
|
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
44 |
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
|
45 |
where |
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
46 |
"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
|
47 |
|
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
48 |
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
|
49 |
where |
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
50 |
"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
|
51 |
|
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
52 |
lemma [code]: |
|
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' = [True, False]" |
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
54 |
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
|
55 |
|
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
56 |
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
|
57 |
where |
|
774752af4a1f
parenthesing let-expressions in OCaml similar to case expressions avoids precendence problems due to ambiguous scope;
haftmann
parents:
59842
diff
changeset
|
58 |
"check_list = (if funny_list = funny_list' then () else undefined)" |
| 31378 | 59 |
|
|
61130
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
61129
diff
changeset
|
60 |
text \<open>Explicit check in @{text Scala} for correct bracketing of abstractions\<close>
|
|
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
61129
diff
changeset
|
61 |
|
|
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
61129
diff
changeset
|
62 |
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
|
63 |
where |
|
8e736ce4c6f4
unconditional parenthesing of (chained) abstractions in Scala, with explicit regression setup
haftmann
parents:
61129
diff
changeset
|
64 |
"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
|
65 |
|
| 21917 | 66 |
end |