src/HOL/Library/Predicate_Compile_Alternative_Defs.thy
author bulwahn
Wed, 31 Mar 2010 16:44:41 +0200
changeset 36053 29e242e9e9a3
parent 35953 0460ff79bb52
child 36246 43fecedff8cf
permissions -rw-r--r--
adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     1
theory Predicate_Compile_Alternative_Defs
36053
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
     2
imports Main
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     3
begin
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     4
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     5
section {* Common constants *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     6
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     7
declare HOL.if_bool_eq_disj[code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     8
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
     9
setup {* Predicate_Compile_Data.ignore_consts [@{const_name Let}] *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    10
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    11
section {* Pairs *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    12
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    13
setup {* Predicate_Compile_Data.ignore_consts [@{const_name fst}, @{const_name snd}, @{const_name split}] *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    14
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    15
section {* Bounded quantifiers *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    16
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    17
declare Ball_def[code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    18
declare Bex_def[code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    19
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    20
section {* Set operations *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    21
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    22
declare Collect_def[code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    23
declare mem_def[code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    24
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    25
declare eq_reflection[OF empty_def, code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    26
declare insert_code[code_pred_def]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    27
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    28
declare subset_iff[code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    29
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    30
declare Int_def[code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    31
declare eq_reflection[OF Un_def, code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    32
declare eq_reflection[OF UNION_def, code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    33
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    34
lemma Diff[code_pred_inline]:
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    35
  "(A - B) = (%x. A x \<and> \<not> B x)"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    36
by (auto simp add: mem_def)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    37
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    38
lemma set_equality[code_pred_inline]:
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    39
  "(A = B) = ((\<forall>x. A x \<longrightarrow> B x) \<and> (\<forall>x. B x \<longrightarrow> A x))"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    40
by (fastsimp simp add: mem_def)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    41
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    42
section {* Setup for Numerals *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    43
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    44
setup {* Predicate_Compile_Data.ignore_consts [@{const_name number_of}] *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    45
setup {* Predicate_Compile_Data.keep_functions [@{const_name number_of}] *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    46
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    47
setup {* Predicate_Compile_Data.ignore_consts [@{const_name div}, @{const_name mod}, @{const_name times}] *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    48
36053
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    49
section {* Arithmetic operations *}
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    50
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    51
subsection {* Arithmetic on naturals and integers *}
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    52
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    53
definition plus_eq_nat :: "nat => nat => nat => bool"
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    54
where
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    55
  "plus_eq_nat x y z = (x + y = z)"
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    56
36053
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    57
definition minus_eq_nat :: "nat => nat => nat => bool"
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    58
where
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    59
  "minus_eq_nat x y z = (x - y = z)"
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    60
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    61
definition plus_eq_int :: "int => int => int => bool"
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    62
where
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    63
  "plus_eq_int x y z = (x + y = z)"
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    64
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    65
definition minus_eq_int :: "int => int => int => bool"
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    66
where
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    67
  "minus_eq_int x y z = (x - y = z)"
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    68
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    69
definition subtract
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    70
where
36053
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    71
  [code_inline]: "subtract x y = y - x"
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    72
36053
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    73
setup {*
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    74
let
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    75
  val Fun = Predicate_Compile_Aux.Fun
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    76
  val Input = Predicate_Compile_Aux.Input
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    77
  val Output = Predicate_Compile_Aux.Output
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    78
  val Bool = Predicate_Compile_Aux.Bool
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    79
  val iio = Fun (Input, Fun (Input, Fun (Output, Bool)))
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    80
  val ioi = Fun (Input, Fun (Output, Fun (Input, Bool)))
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    81
  val oii = Fun (Output, Fun (Input, Fun (Input, Bool)))
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    82
  val ooi = Fun (Output, Fun (Output, Fun (Input, Bool)))
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    83
  val plus_nat = Predicate_Compile_Core.functional_compilation @{const_name plus} iio
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    84
  val minus_nat = Predicate_Compile_Core.functional_compilation @{const_name "minus"} iio
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    85
  fun subtract_nat compfuns (_ : typ) =
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    86
    let
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    87
      val T = Predicate_Compile_Aux.mk_predT compfuns @{typ nat}
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    88
    in
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    89
      absdummy (@{typ nat}, absdummy (@{typ nat},
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    90
        Const (@{const_name "If"}, @{typ bool} --> T --> T --> T) $
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    91
          (@{term "op > :: nat => nat => bool"} $ Bound 1 $ Bound 0) $
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    92
          Predicate_Compile_Aux.mk_bot compfuns @{typ nat} $
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    93
          Predicate_Compile_Aux.mk_single compfuns
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    94
          (@{term "op - :: nat => nat => nat"} $ Bound 0 $ Bound 1)))
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    95
    end
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    96
  fun enumerate_addups_nat compfuns (_ : typ) =
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    97
    absdummy (@{typ nat}, Predicate_Compile_Aux.mk_iterate_upto compfuns @{typ "nat * nat"}
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    98
    (absdummy (@{typ code_numeral}, @{term "Pair :: nat => nat => nat * nat"} $
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
    99
      (@{term "Code_Numeral.nat_of"} $ Bound 0) $
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   100
      (@{term "op - :: nat => nat => nat"} $ Bound 1 $ (@{term "Code_Numeral.nat_of"} $ Bound 0))),
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   101
      @{term "0 :: code_numeral"}, @{term "Code_Numeral.of_nat"} $ Bound 0))
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   102
  fun enumerate_nats compfuns  (_ : typ) =
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   103
    let
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   104
      val (single_const, _) = strip_comb (Predicate_Compile_Aux.mk_single compfuns @{term "0 :: nat"})
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   105
      val T = Predicate_Compile_Aux.mk_predT compfuns @{typ nat}
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   106
    in
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   107
      absdummy(@{typ nat}, absdummy (@{typ nat},
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   108
        Const (@{const_name If}, @{typ bool} --> T --> T --> T) $
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   109
          (@{term "op = :: nat => nat => bool"} $ Bound 0 $ @{term "0::nat"}) $
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   110
          (Predicate_Compile_Aux.mk_iterate_upto compfuns @{typ nat} (@{term "Code_Numeral.nat_of"},
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   111
            @{term "0::code_numeral"}, @{term "Code_Numeral.of_nat"} $ Bound 1)) $
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   112
            (single_const $ (@{term "op + :: nat => nat => nat"} $ Bound 1 $ Bound 0))))
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   113
    end
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   114
in
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   115
  Predicate_Compile_Core.force_modes_and_compilations @{const_name plus_eq_nat}
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   116
    [(iio, (plus_nat, false)), (oii, (subtract_nat, false)), (ioi, (subtract_nat, false)),
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   117
     (ooi, (enumerate_addups_nat, false))]
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   118
  #> Predicate_Compile_Fun.add_function_predicate_translation
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   119
       (@{term "plus :: nat => nat => nat"}, @{term "plus_eq_nat"})
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   120
  #> Predicate_Compile_Core.force_modes_and_compilations @{const_name minus_eq_nat}
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   121
       [(iio, (minus_nat, false)), (oii, (enumerate_nats, false))]
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   122
  #> Predicate_Compile_Fun.add_function_predicate_translation
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   123
      (@{term "minus :: nat => nat => nat"}, @{term "minus_eq_nat"})
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   124
  #> Predicate_Compile_Core.force_modes_and_functions @{const_name plus_eq_int}
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   125
    [(iio, (@{const_name plus}, false)), (ioi, (@{const_name subtract}, false)),
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   126
     (oii, (@{const_name subtract}, false))]
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   127
  #> Predicate_Compile_Fun.add_function_predicate_translation
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   128
       (@{term "plus :: int => int => int"}, @{term "plus_eq_int"})
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   129
  #> Predicate_Compile_Core.force_modes_and_functions @{const_name minus_eq_int}
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   130
    [(iio, (@{const_name minus}, false)), (oii, (@{const_name plus}, false)),
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   131
     (ioi, (@{const_name minus}, false))]
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   132
  #> Predicate_Compile_Fun.add_function_predicate_translation
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   133
      (@{term "minus :: int => int => int"}, @{term "minus_eq_int"})
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   134
end
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   135
*}
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   136
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   137
subsection {* Inductive definitions for ordering on naturals *}
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   138
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   139
inductive less_nat
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   140
where
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   141
  "less_nat 0 (Suc y)"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   142
| "less_nat x y ==> less_nat (Suc x) (Suc y)"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   143
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   144
lemma [code_pred_inline]:
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   145
  "x < y = less_nat x y"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   146
apply (rule iffI)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   147
apply (induct x arbitrary: y)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   148
apply (case_tac y) apply (auto intro: less_nat.intros)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   149
apply (case_tac y)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   150
apply (auto intro: less_nat.intros)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   151
apply (induct rule: less_nat.induct)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   152
apply auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   153
done
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   154
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   155
inductive less_eq_nat
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   156
where
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   157
  "less_eq_nat 0 y"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   158
| "less_eq_nat x y ==> less_eq_nat (Suc x) (Suc y)"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   159
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   160
lemma [code_pred_inline]:
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   161
"x <= y = less_eq_nat x y"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   162
apply (rule iffI)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   163
apply (induct x arbitrary: y)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   164
apply (auto intro: less_eq_nat.intros)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   165
apply (case_tac y) apply (auto intro: less_eq_nat.intros)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   166
apply (induct rule: less_eq_nat.induct)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   167
apply auto done
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   168
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   169
section {* Alternative list definitions *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   170
36053
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   171
subsection {* Alternative rules for length *}
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   172
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   173
definition size_list :: "'a list => nat"
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   174
where "size_list = size"
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   175
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   176
lemma size_list_simps:
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   177
  "size_list [] = 0"
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   178
  "size_list (x # xs) = Suc (size_list xs)"
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   179
by (auto simp add: size_list_def)
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   180
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   181
declare size_list_simps[code_pred_def]
29e242e9e9a3 adding setup for handling arithmetic of natural numbers and integers precisely and more efficiently in the predicate compiler; changed alternative size_list definition for predicate compiler
bulwahn
parents: 35953
diff changeset
   182
declare size_list_def[symmetric, code_pred_inline]
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   183
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   184
subsection {* Alternative rules for set *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   185
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   186
lemma set_ConsI1 [code_pred_intro]:
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   187
  "set (x # xs) x"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   188
unfolding mem_def[symmetric, of _ x]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   189
by auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   190
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   191
lemma set_ConsI2 [code_pred_intro]:
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   192
  "set xs x ==> set (x' # xs) x" 
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   193
unfolding mem_def[symmetric, of _ x]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   194
by auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   195
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   196
code_pred [skip_proof] set
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   197
proof -
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   198
  case set
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   199
  from this show thesis
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   200
    apply (case_tac xb)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   201
    apply auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   202
    unfolding mem_def[symmetric, of _ xc]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   203
    apply auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   204
    unfolding mem_def
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   205
    apply fastsimp
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   206
    done
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   207
qed
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   208
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   209
subsection {* Alternative rules for list_all2 *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   210
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   211
lemma list_all2_NilI [code_pred_intro]: "list_all2 P [] []"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   212
by auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   213
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   214
lemma list_all2_ConsI [code_pred_intro]: "list_all2 P xs ys ==> P x y ==> list_all2 P (x#xs) (y#ys)"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   215
by auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   216
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   217
code_pred [skip_proof] list_all2
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   218
proof -
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   219
  case list_all2
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   220
  from this show thesis
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   221
    apply -
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   222
    apply (case_tac xb)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   223
    apply (case_tac xc)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   224
    apply auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   225
    apply (case_tac xc)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   226
    apply auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   227
    apply fastsimp
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   228
    done
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   229
qed
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   230
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   231
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   232
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   233
end