src/HOL/Codegenerator_Test/RBT_Set_Test.thy
author blanchet
Fri, 15 Feb 2013 09:17:20 +0100
changeset 51133 fb16c4276620
parent 51116 0dac0158b8d4
child 51139 c8e3cf3520b3
permissions -rw-r--r--
tuning
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48624
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
     1
(*  Title:      HOL/Codegenerator_Test/RBT_Set_Test.thy
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
     2
    Author:     Ondrej Kuncar
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
     3
*)
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
     4
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
     5
header {* Test of the code generator using an implementation of sets by RBT trees *}
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
     6
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
     7
theory RBT_Set_Test
49948
744934b818c7 moved quite generic material from theory Enum to more appropriate places
haftmann
parents: 48624
diff changeset
     8
imports "~~/src/HOL/Library/Library" "~~/src/HOL/Library/RBT_Set"
48624
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
     9
begin
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    10
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    11
(* 
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    12
   The following code equations have to be deleted because they use 
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    13
   lists to implement sets in the code generetor. 
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    14
*)
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    15
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    16
lemma [code, code del]:
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    17
  "Sup_pred_inst.Sup_pred = Sup_pred_inst.Sup_pred" ..
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    18
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    19
lemma [code, code del]:
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    20
  "Inf_pred_inst.Inf_pred = Inf_pred_inst.Inf_pred" ..
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    21
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    22
lemma [code, code del]:
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    23
  "pred_of_set = pred_of_set" ..
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    24
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    25
lemma [code, code del]:
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    26
  "acc = acc" ..
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    27
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
    28
lemmas [code del] =
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
    29
  finite_set_code finite_coset_code 
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
    30
  equal_set_code
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
    31
48624
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    32
(*
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    33
  If the code generation ends with an exception with the following message:
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    34
  '"List.set" is not a constructor, on left hand side of equation: ...',
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    35
  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
    36
  or implemented for RBT trees.
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    37
*)
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    38
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
    39
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
    40
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
    41
(* 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
    42
(* 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
    43
   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
    44
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
    45
class ccpo_linorder = ccpo + linorder
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
    46
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
    47
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
    48
  "(Complete_Partial_Order.admissible :: ('a :: ccpo_linorder \<Rightarrow> bool) \<Rightarrow> bool) P = 
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
    49
    (\<forall>A. Complete_Partial_Order.chain (op \<le>) A \<longrightarrow> (\<forall>x\<in>A. P x) \<longrightarrow> P (Sup A))"
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
    50
unfolding admissible_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
    51
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
    52
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
    53
  "(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
    54
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
    55
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
    56
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
    57
  "(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
    58
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
    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: 49948
diff changeset
    60
export_code _ checking Scala?
48624
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    61
9b71daba4ec7 add testing file for RBT_Set
kuncar
parents:
diff changeset
    62
end