| author | blanchet | 
| Sun, 12 Oct 2014 21:52:45 +0200 | |
| changeset 58654 | 3e1cad27fc2f | 
| parent 54437 | 0060957404c7 | 
| child 58889 | 5b7a9633cfa8 | 
| 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 | |
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51139diff
changeset | 4 | header {* Pervasive test of code generator *}
 | 
| 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 | ||
| 14 | (* | |
| 15 | The following code equations have to be deleted because they use | |
| 16 | lists to implement sets in the code generetor. | |
| 17 | *) | |
| 18 | ||
| 19 | lemma [code, code del]: | |
| 20 | "Sup_pred_inst.Sup_pred = Sup_pred_inst.Sup_pred" .. | |
| 21 | ||
| 22 | lemma [code, code del]: | |
| 23 | "Inf_pred_inst.Inf_pred = Inf_pred_inst.Inf_pred" .. | |
| 24 | ||
| 25 | lemma [code, code del]: | |
| 26 | "pred_of_set = pred_of_set" .. | |
| 27 | ||
| 28 | lemma [code, code del]: | |
| 54295 | 29 | "Wellfounded.acc = Wellfounded.acc" .. | 
| 48624 | 30 | |
| 51139 
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
 Andreas Lochbihler parents: 
51116diff
changeset | 31 | lemma [code, code del]: | 
| 
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
 Andreas Lochbihler parents: 
51116diff
changeset | 32 | "Cardinality.card' = Cardinality.card'" .. | 
| 
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
 Andreas Lochbihler parents: 
51116diff
changeset | 33 | |
| 
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
 Andreas Lochbihler parents: 
51116diff
changeset | 34 | lemma [code, code del]: | 
| 
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
 Andreas Lochbihler parents: 
51116diff
changeset | 35 | "Cardinality.finite' = Cardinality.finite'" .. | 
| 
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
 Andreas Lochbihler parents: 
51116diff
changeset | 36 | |
| 
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
 Andreas Lochbihler parents: 
51116diff
changeset | 37 | lemma [code, code del]: | 
| 
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
 Andreas Lochbihler parents: 
51116diff
changeset | 38 | "Cardinality.subset' = Cardinality.subset'" .. | 
| 
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
 Andreas Lochbihler parents: 
51116diff
changeset | 39 | |
| 
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
 Andreas Lochbihler parents: 
51116diff
changeset | 40 | lemma [code, code del]: | 
| 
c8e3cf3520b3
partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
 Andreas Lochbihler parents: 
51116diff
changeset | 41 | "Cardinality.eq_set = Cardinality.eq_set" .. | 
| 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 | 42 | |
| 54437 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 43 | lemma [code, code del]: | 
| 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 44 | "(Gcd :: nat set \<Rightarrow> nat) = Gcd" .. | 
| 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 45 | |
| 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 46 | lemma [code, code del]: | 
| 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 47 | "(Lcm :: nat set \<Rightarrow> nat) = Lcm" .. | 
| 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 48 | |
| 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 49 | lemma [code, code del]: | 
| 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 50 | "(Gcd :: int set \<Rightarrow> int) = Gcd" .. | 
| 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 51 | |
| 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 52 | lemma [code, code del]: | 
| 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 53 | "(Lcm :: int set \<Rightarrow> int) = Lcm" .. | 
| 
0060957404c7
proper code equations for Gcd and Lcm on nat and int
 haftmann parents: 
54295diff
changeset | 54 | |
| 48624 | 55 | (* | 
| 56 | If the code generation ends with an exception with the following message: | |
| 57 | '"List.set" is not a constructor, on left hand side of equation: ...', | |
| 58 | the code equation in question has to be either deleted (like many others in this file) | |
| 59 | or implemented for RBT trees. | |
| 60 | *) | |
| 61 | ||
| 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 | 62 | 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 | 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 | (* 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 | 65 | (* 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 | 66 | 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 | 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 |   "(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 | 70 | 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 | 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 | 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 | 73 |   "(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 | 74 | 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 | 75 | |
| 
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 | 76 | export_code _ checking Scala? | 
| 48624 | 77 | |
| 78 | end |