src/HOL/Codegenerator_Test/RBT_Set_Test.thy
changeset 51161 6ed12ae3b3e1
parent 51160 599ff65b85e2
child 51162 310b94ed1815
equal deleted inserted replaced
51160:599ff65b85e2 51161:6ed12ae3b3e1
     1 (*  Title:      HOL/Codegenerator_Test/RBT_Set_Test.thy
       
     2     Author:     Ondrej Kuncar
       
     3 *)
       
     4 
       
     5 header {* Test of the code generator using an implementation of sets by RBT trees *}
       
     6 
       
     7 theory RBT_Set_Test
       
     8 imports "~~/src/HOL/Library/Library" "~~/src/HOL/Library/RBT_Set"
       
     9 begin
       
    10 
       
    11 (* 
       
    12    The following code equations have to be deleted because they use 
       
    13    lists to implement sets in the code generetor. 
       
    14 *)
       
    15 
       
    16 lemma [code, code del]:
       
    17   "Sup_pred_inst.Sup_pred = Sup_pred_inst.Sup_pred" ..
       
    18 
       
    19 lemma [code, code del]:
       
    20   "Inf_pred_inst.Inf_pred = Inf_pred_inst.Inf_pred" ..
       
    21 
       
    22 lemma [code, code del]:
       
    23   "pred_of_set = pred_of_set" ..
       
    24 
       
    25 lemma [code, code del]:
       
    26   "acc = acc" ..
       
    27 
       
    28 lemma [code, code del]:
       
    29   "Cardinality.card' = Cardinality.card'" ..
       
    30 
       
    31 lemma [code, code del]:
       
    32   "Cardinality.finite' = Cardinality.finite'" ..
       
    33 
       
    34 lemma [code, code del]:
       
    35   "Cardinality.subset' = Cardinality.subset'" ..
       
    36 
       
    37 lemma [code, code del]:
       
    38   "Cardinality.eq_set = Cardinality.eq_set" ..
       
    39 
       
    40 (*
       
    41   If the code generation ends with an exception with the following message:
       
    42   '"List.set" is not a constructor, on left hand side of equation: ...',
       
    43   the code equation in question has to be either deleted (like many others in this file) 
       
    44   or implemented for RBT trees.
       
    45 *)
       
    46 
       
    47 export_code _ checking SML OCaml? Haskell?
       
    48 
       
    49 (* Extra setup to make Scala happy *)
       
    50 (* If the compilation fails with an error "ambiguous implicit values",
       
    51    read \<section>7.1 in the Code Generation Manual *)
       
    52 
       
    53 class ccpo_linorder = ccpo + linorder
       
    54 
       
    55 lemma [code]:
       
    56   "(Complete_Partial_Order.admissible :: ('a :: ccpo_linorder \<Rightarrow> bool) \<Rightarrow> bool) P = 
       
    57     (\<forall>A. Complete_Partial_Order.chain (op \<le>) A \<longrightarrow> (\<forall>x\<in>A. P x) \<longrightarrow> P (Sup A))"
       
    58 unfolding admissible_def by blast
       
    59 
       
    60 lemma [code]:
       
    61   "(gfp :: ('a :: complete_linorder \<Rightarrow> 'a) \<Rightarrow> 'a) f = Sup {u. u \<le> f u}"
       
    62 unfolding gfp_def by blast
       
    63 
       
    64 lemma [code]:
       
    65   "(lfp :: ('a :: complete_linorder \<Rightarrow> 'a) \<Rightarrow> 'a) f = Inf {u. f u \<le> u}"
       
    66 unfolding lfp_def by blast
       
    67 
       
    68 export_code _ checking Scala?
       
    69 
       
    70 end