src/HOL/Codegenerator_Test/Generate_Efficient_Datastructures.thy
author wenzelm
Mon, 24 Oct 2016 14:05:22 +0200
changeset 64371 213cf4215b40
parent 63965 d510b816ea41
child 64786 340db65fd2c1
permissions -rw-r--r--
more robust;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 51139
diff changeset
     1
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 51139
diff 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: 51139
diff changeset
     3
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63134
diff changeset
     4
section \<open>Pervasive test of code generator\<close>
48624
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
     5
51161
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 51139
diff changeset
     6
theory Generate_Efficient_Datastructures
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 51139
diff changeset
     7
imports
6ed12ae3b3e1 attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents: 51139
diff changeset
     8
  Candidates
51599
1559e9266280 optionalized very specific code setup for multisets
haftmann
parents: 51161
diff 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: 51139
diff 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: 51139
diff changeset
    11
  "~~/src/HOL/Library/RBT_Set"
48624
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    12
begin
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    13
59484
a130ae7a9398 slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents: 58889
diff changeset
    14
setup \<open>
59842
wenzelm
parents: 59484
diff changeset
    15
fn thy =>
59484
a130ae7a9398 slightly more standard code setup for String.literal, with explicit special case in predicate compiler
haftmann
parents: 58889
diff changeset
    16
let
59842
wenzelm
parents: 59484
diff changeset
    17
  val tycos = Sign.logical_types thy;
wenzelm
parents: 59484
diff changeset
    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: 58889
diff changeset
    19
    @{const_name "Quickcheck_Narrowing.partial_term_of"})) tycos;
59842
wenzelm
parents: 59484
diff changeset
    20
in fold Code.del_eqns consts thy end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63134
diff changeset
    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: 58889
diff changeset
    22
48624
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    23
(* 
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    24
   The following code equations have to be deleted because they use 
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    25
   lists to implement sets in the code generetor. 
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    26
*)
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    27
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    28
lemma [code, code del]:
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    29
  "Sup_pred_inst.Sup_pred = Sup_pred_inst.Sup_pred" ..
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    30
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    31
lemma [code, code del]:
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    32
  "Inf_pred_inst.Inf_pred = Inf_pred_inst.Inf_pred" ..
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    33
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    34
lemma [code, code del]:
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    35
  "pred_of_set = pred_of_set" ..
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    36
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    37
lemma [code, code del]:
54295
45a5523d4a63 qualifed popular user space names
haftmann
parents: 53361
diff changeset
    38
  "Wellfounded.acc = Wellfounded.acc" ..
48624
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    39
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: 51116
diff 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: 51116
diff changeset
    41
  "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: 51116
diff changeset
    42
c8e3cf3520b3 partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents: 51116
diff changeset
    43
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: 51116
diff changeset
    44
  "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: 51116
diff changeset
    45
c8e3cf3520b3 partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents: 51116
diff changeset
    46
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: 51116
diff changeset
    47
  "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: 51116
diff changeset
    48
c8e3cf3520b3 partially revert 0dac0158b8d4 as it too aggressively spreads card_UNIV type class whose dictionary constructions can slow down generated code;
Andreas Lochbihler
parents: 51116
diff changeset
    49
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: 51116
diff changeset
    50
  "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: 50996
diff changeset
    51
54437
0060957404c7 proper code equations for Gcd and Lcm on nat and int
haftmann
parents: 54295
diff changeset
    52
lemma [code, code del]:
0060957404c7 proper code equations for Gcd and Lcm on nat and int
haftmann
parents: 54295
diff changeset
    53
  "(Gcd :: nat set \<Rightarrow> nat) = Gcd" ..
0060957404c7 proper code equations for Gcd and Lcm on nat and int
haftmann
parents: 54295
diff changeset
    54
0060957404c7 proper code equations for Gcd and Lcm on nat and int
haftmann
parents: 54295
diff changeset
    55
lemma [code, code del]:
0060957404c7 proper code equations for Gcd and Lcm on nat and int
haftmann
parents: 54295
diff changeset
    56
  "(Lcm :: nat set \<Rightarrow> nat) = Lcm" ..
0060957404c7 proper code equations for Gcd and Lcm on nat and int
haftmann
parents: 54295
diff changeset
    57
0060957404c7 proper code equations for Gcd and Lcm on nat and int
haftmann
parents: 54295
diff changeset
    58
lemma [code, code del]:
0060957404c7 proper code equations for Gcd and Lcm on nat and int
haftmann
parents: 54295
diff changeset
    59
  "(Gcd :: int set \<Rightarrow> int) = Gcd" ..
0060957404c7 proper code equations for Gcd and Lcm on nat and int
haftmann
parents: 54295
diff changeset
    60
0060957404c7 proper code equations for Gcd and Lcm on nat and int
haftmann
parents: 54295
diff changeset
    61
lemma [code, code del]:
0060957404c7 proper code equations for Gcd and Lcm on nat and int
haftmann
parents: 54295
diff changeset
    62
  "(Lcm :: int set \<Rightarrow> int) = Lcm" ..
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    63
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    64
lemma [code, code del]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    65
  "(Gcd :: _ poly set \<Rightarrow> _) = Gcd" ..
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    66
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    67
lemma [code, code del]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    68
  "(Lcm :: _ poly set \<Rightarrow> _) = Lcm" ..
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    69
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    70
lemma [code, code del]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    71
  "Gcd_eucl = Gcd_eucl" ..
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    72
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    73
lemma [code, code del]:
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    74
  "Lcm_eucl = Lcm_eucl" ..
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 59842
diff changeset
    75
63134
aa573306a9cd Removed problematic code equation for set_permutations
eberlm
parents: 63133
diff changeset
    76
lemma [code, code del]:
aa573306a9cd Removed problematic code equation for set_permutations
eberlm
parents: 63133
diff changeset
    77
  "permutations_of_set = permutations_of_set" ..
aa573306a9cd Removed problematic code equation for set_permutations
eberlm
parents: 63133
diff changeset
    78
63965
d510b816ea41 Set_Permutations replaced by more general Multiset_Permutations
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
    79
lemma [code, code del]:
d510b816ea41 Set_Permutations replaced by more general Multiset_Permutations
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
    80
  "permutations_of_multiset = permutations_of_multiset" ..
d510b816ea41 Set_Permutations replaced by more general Multiset_Permutations
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
    81
48624
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    82
(*
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    83
  If the code generation ends with an exception with the following message:
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    84
  '"List.set" is not a constructor, on left hand side of equation: ...',
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    85
  the code equation in question has to be either deleted (like many others in this file) 
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    86
  or implemented for RBT trees.
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    87
*)
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    88
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: 49948
diff changeset
    89
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: 49948
diff changeset
    90
51ad7b4ac096 restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
kuncar
parents: 49948
diff changeset
    91
(* 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: 49948
diff changeset
    92
(* 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: 49948
diff changeset
    93
   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: 49948
diff changeset
    94
51ad7b4ac096 restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
kuncar
parents: 49948
diff changeset
    95
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: 49948
diff changeset
    96
  "(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: 49948
diff changeset
    97
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: 49948
diff changeset
    98
51ad7b4ac096 restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
kuncar
parents: 49948
diff changeset
    99
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: 49948
diff changeset
   100
  "(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: 49948
diff changeset
   101
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: 49948
diff changeset
   102
51ad7b4ac096 restore code equations for List.set in RBT_Set; make Scala happy according to 7.1 in the code generator manual
kuncar
parents: 49948
diff changeset
   103
export_code _ checking Scala?
48624
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
   104
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
   105
end