src/HOL/Codegenerator_Test/Generate_Efficient_Datastructures.thy
author wenzelm
Tue, 19 Dec 2017 13:58:12 +0100
changeset 67226 ec32cdaab97b
parent 66453 cc19f7ca2ed6
child 73477 1d8a79aa2a99
permissions -rw-r--r--
isabelle update_cartouches -c -t;
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
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66404
diff changeset
     9
  "HOL-Library.DAList_Multiset"
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66404
diff changeset
    10
  "HOL-Library.RBT_Mapping"
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66404
diff changeset
    11
  "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
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    14
text \<open>
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff 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: 63965
diff 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: 63965
diff 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: 50996
diff changeset
    18
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    19
declare [[code drop:
66016
haftmann
parents: 66014
diff changeset
    20
  "Inf :: _ Predicate.pred set \<Rightarrow> _"
haftmann
parents: 66014
diff changeset
    21
  "Sup :: _ Predicate.pred set \<Rightarrow> _"
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    22
  pred_of_set
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    23
  Wellfounded.acc
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    24
  Cardinality.card'
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    25
  Cardinality.finite'
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    26
  Cardinality.subset'
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    27
  Cardinality.eq_set
66016
haftmann
parents: 66014
diff changeset
    28
  Euclidean_Algorithm.Gcd
haftmann
parents: 66014
diff changeset
    29
  Euclidean_Algorithm.Lcm
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    30
  "Gcd :: _ poly set \<Rightarrow> _"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    31
  "Lcm :: _ poly set \<Rightarrow> _"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    32
  permutations_of_set
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    33
  permutations_of_multiset
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 63965
diff changeset
    34
]]
63965
d510b816ea41 Set_Permutations replaced by more general Multiset_Permutations
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
    35
66016
haftmann
parents: 66014
diff changeset
    36
text \<open>
haftmann
parents: 66014
diff changeset
    37
  If code generation fails with a message like
67226
ec32cdaab97b isabelle update_cartouches -c -t;
wenzelm
parents: 66453
diff changeset
    38
  \<open>"List.set" is not a constructor, on left hand side of equation: ...\<close>,
66016
haftmann
parents: 66014
diff changeset
    39
  feel free to add an RBT implementation for the corrsponding operation
haftmann
parents: 66014
diff changeset
    40
  of delete that code equation (see above).
haftmann
parents: 66014
diff changeset
    41
\<close>
haftmann
parents: 66014
diff changeset
    42
 
65919
b6d458915f1b obsolete special case
haftmann
parents: 64850
diff changeset
    43
export_code _ checking SML OCaml? Haskell? Scala
48624
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    44
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    45
end