| author | paulson <lp15@cam.ac.uk> | 
| Tue, 11 Sep 2018 16:21:54 +0100 | |
| changeset 68975 | 5ce4d117cea7 | 
| parent 67226 | ec32cdaab97b | 
| child 73477 | 1d8a79aa2a99 | 
| permissions | -rw-r--r-- | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51139diff
changeset | 1 | |
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51139diff
changeset | 2 | (* Author: Ondrej Kuncar, TU Muenchen *) | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51139diff
changeset | 3 | |
| 63167 | 4 | section \<open>Pervasive test of code generator\<close> | 
| 48624 | 5 | |
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51139diff
changeset | 6 | theory Generate_Efficient_Datastructures | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51139diff
changeset | 7 | imports | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51139diff
changeset | 8 | Candidates | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66404diff
changeset | 9 | "HOL-Library.DAList_Multiset" | 
| 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66404diff
changeset | 10 | "HOL-Library.RBT_Mapping" | 
| 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66404diff
changeset | 11 | "HOL-Library.RBT_Set" | 
| 48624 | 12 | begin | 
| 13 | ||
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 14 | text \<open> | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 15 | The following code equations have to be deleted because they use | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 16 | lists to implement sets in the code generetor. | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 17 | \<close> | 
| 51116 
0dac0158b8d4
implement code generation for finite, card, op = and op <= for sets always via finite_UNIV and card_UNIV, as fragile rewrites based on sorts are hard to find and debug
 Andreas Lochbihler parents: 
50996diff
changeset | 18 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 19 | declare [[code drop: | 
| 66016 | 20 | "Inf :: _ Predicate.pred set \<Rightarrow> _" | 
| 21 | "Sup :: _ Predicate.pred set \<Rightarrow> _" | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 22 | pred_of_set | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 23 | Wellfounded.acc | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 24 | Cardinality.card' | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 25 | Cardinality.finite' | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 26 | Cardinality.subset' | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 27 | Cardinality.eq_set | 
| 66016 | 28 | Euclidean_Algorithm.Gcd | 
| 29 | Euclidean_Algorithm.Lcm | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 30 | "Gcd :: _ poly set \<Rightarrow> _" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 31 | "Lcm :: _ poly set \<Rightarrow> _" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 32 | permutations_of_set | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 33 | permutations_of_multiset | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 34 | ]] | 
| 63965 
d510b816ea41
Set_Permutations replaced by more general Multiset_Permutations
 eberlm <eberlm@in.tum.de> parents: 
63167diff
changeset | 35 | |
| 66016 | 36 | text \<open> | 
| 37 | If code generation fails with a message like | |
| 67226 | 38 | \<open>"List.set" is not a constructor, on left hand side of equation: ...\<close>, | 
| 66016 | 39 | feel free to add an RBT implementation for the corrsponding operation | 
| 40 | of delete that code equation (see above). | |
| 41 | \<close> | |
| 42 | ||
| 65919 | 43 | export_code _ checking SML OCaml? Haskell? Scala | 
| 48624 | 44 | |
| 45 | end |