| author | wenzelm | 
| Mon, 01 May 2017 13:03:56 +0200 | |
| changeset 65666 | 45d0692bb019 | 
| parent 64850 | fc9265882329 | 
| child 65919 | b6d458915f1b | 
| 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 | 
| 51599 
1559e9266280
optionalized very specific code setup for multisets
 haftmann parents: 
51161diff
changeset | 9 | "~~/src/HOL/Library/DAList_Multiset" | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51139diff
changeset | 10 | "~~/src/HOL/Library/RBT_Mapping" | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51139diff
changeset | 11 | "~~/src/HOL/Library/RBT_Set" | 
| 48624 | 12 | begin | 
| 13 | ||
| 59484 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
58889diff
changeset | 14 | setup \<open> | 
| 59842 | 15 | fn thy => | 
| 59484 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
58889diff
changeset | 16 | let | 
| 59842 | 17 | val tycos = Sign.logical_types thy; | 
| 18 | 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: 
58889diff
changeset | 19 |     @{const_name "Quickcheck_Narrowing.partial_term_of"})) tycos;
 | 
| 59842 | 20 | in fold Code.del_eqns consts thy end | 
| 63167 | 21 | \<close> \<comment> \<open>drop technical stuff from \<open>Quickcheck_Narrowing\<close> which is tailored towards Haskell\<close> | 
| 59484 
a130ae7a9398
slightly more standard code setup for String.literal, with explicit special case in predicate compiler
 haftmann parents: 
58889diff
changeset | 22 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 23 | text \<open> | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 24 | 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 | 25 | 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 | 26 | \<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 | 27 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 28 | declare [[code drop: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 29 | Sup_pred_inst.Sup_pred | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 30 | Inf_pred_inst.Inf_pred | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 31 | pred_of_set | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 32 | Wellfounded.acc | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 33 | Cardinality.card' | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 34 | Cardinality.finite' | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 35 | Cardinality.subset' | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 36 | Cardinality.eq_set | 
| 64850 | 37 | Gcd_fin | 
| 38 | Lcm_fin | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 39 | "Gcd :: nat set \<Rightarrow> nat" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 40 | "Lcm :: nat set \<Rightarrow> nat" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 41 | "Gcd :: int set \<Rightarrow> int" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 42 | "Lcm :: int set \<Rightarrow> int" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 43 | "Gcd :: _ poly set \<Rightarrow> _" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 44 | "Lcm :: _ poly set \<Rightarrow> _" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 45 | Euclidean_Algorithm.Gcd | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 46 | Euclidean_Algorithm.Lcm | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 47 | permutations_of_set | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 48 | permutations_of_multiset | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
63965diff
changeset | 49 | ]] | 
| 63965 
d510b816ea41
Set_Permutations replaced by more general Multiset_Permutations
 eberlm <eberlm@in.tum.de> parents: 
63167diff
changeset | 50 | |
| 48624 | 51 | (* | 
| 52 | If the code generation ends with an exception with the following message: | |
| 53 | '"List.set" is not a constructor, on left hand side of equation: ...', | |
| 54 | the code equation in question has to be either deleted (like many others in this file) | |
| 55 | or implemented for RBT trees. | |
| 56 | *) | |
| 57 | ||
| 50996 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 58 | export_code _ checking SML OCaml? Haskell? | 
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 59 | |
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 60 | (* Extra setup to make Scala happy *) | 
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 61 | (* If the compilation fails with an error "ambiguous implicit values", | 
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 62 | read \<section>7.1 in the Code Generation Manual *) | 
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 63 | |
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 64 | lemma [code]: | 
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 65 |   "(gfp :: ('a :: complete_linorder \<Rightarrow> 'a) \<Rightarrow> 'a) f = Sup {u. u \<le> f u}"
 | 
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 66 | unfolding gfp_def by blast | 
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 67 | |
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 68 | lemma [code]: | 
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 69 |   "(lfp :: ('a :: complete_linorder \<Rightarrow> 'a) \<Rightarrow> 'a) f = Inf {u. f u \<le> u}"
 | 
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 70 | unfolding lfp_def by blast | 
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 71 | |
| 
51ad7b4ac096
restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
 kuncar parents: 
49948diff
changeset | 72 | export_code _ checking Scala? | 
| 48624 | 73 | |
| 74 | end |