| author | hoelzl | 
| Mon, 24 Nov 2014 12:20:14 +0100 | |
| changeset 59048 | 7dc8ac6f0895 | 
| parent 58889 | 5b7a9633cfa8 | 
| child 59484 | a130ae7a9398 | 
| 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: 
37407diff
changeset | 6 | theory Candidates | 
| 21917 | 7 | imports | 
| 27421 | 8 | Complex_Main | 
| 51160 
599ff65b85e2
systematic conversions between nat and nibble/char;
 haftmann parents: 
49077diff
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: 
51160diff
changeset | 10 | "~~/src/HOL/Library/Sublist_Order" | 
| 51173 | 11 | "~~/src/HOL/Number_Theory/Eratosthenes" | 
| 35303 | 12 | "~~/src/HOL/ex/Records" | 
| 21917 | 13 | begin | 
| 14 | ||
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51160diff
changeset | 15 | 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: 
51160diff
changeset | 16 | where | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51160diff
changeset | 17 | empty: "sublist [] xs" | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51160diff
changeset | 18 | | 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: 
51160diff
changeset | 19 | | take: "sublist ys xs \<Longrightarrow> sublist (x # ys) (x # xs)" | 
| 33500 | 20 | |
| 21 | code_pred sublist . | |
| 22 | ||
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51160diff
changeset | 23 | code_reserved SML upto -- {* avoid popular infix *}
 | 
| 31378 | 24 | |
| 21917 | 25 | end |