src/HOL/Library/Predicate_Compile_Alternative_Defs.thy
author blanchet
Wed, 28 May 2014 17:42:36 +0200
changeset 57108 dc0b4f50e288
parent 56679 5545bfdfefcc
child 60045 cd2b6debac18
permissions -rw-r--r--
more generous max number of suggestions, for more safety
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
36253
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36246
diff changeset
     9
declare bool_diff_def[code_pred_inline]
46905
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46884
diff changeset
    10
declare inf_bool_def[abs_def, code_pred_inline]
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46884
diff changeset
    11
declare less_bool_def[abs_def, code_pred_inline]
6b1c0a80a57a prefer abs_def over def_raw;
wenzelm
parents: 46884
diff changeset
    12
declare le_bool_def[abs_def, code_pred_inline]
36253
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36246
diff changeset
    13
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36246
diff changeset
    14
lemma min_bool_eq [code_pred_inline]: "(min :: bool => bool => bool) == (op &)"
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 45461
diff changeset
    15
by (rule eq_reflection) (auto simp add: fun_eq_iff min_def)
36253
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36246
diff changeset
    16
39650
2a35950ec495 handling equivalences smarter in the predicate compiler
bulwahn
parents: 39302
diff changeset
    17
lemma [code_pred_inline]: 
2a35950ec495 handling equivalences smarter in the predicate compiler
bulwahn
parents: 39302
diff changeset
    18
  "((A::bool) ~= (B::bool)) = ((A & ~ B) | (B & ~ A))"
2a35950ec495 handling equivalences smarter in the predicate compiler
bulwahn
parents: 39302
diff changeset
    19
by fast
2a35950ec495 handling equivalences smarter in the predicate compiler
bulwahn
parents: 39302
diff changeset
    20
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    21
setup {* Predicate_Compile_Data.ignore_consts [@{const_name Let}] *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    22
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    23
section {* Pairs *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    24
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 54489
diff changeset
    25
setup {* Predicate_Compile_Data.ignore_consts [@{const_name fst}, @{const_name snd}, @{const_name case_prod}] *}
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    26
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    27
section {* Bounded quantifiers *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    28
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    29
declare Ball_def[code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    30
declare Bex_def[code_pred_inline]
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    31
47840
732ea1f08e3f removing obsolete setup for sets now that sets are executable
bulwahn
parents: 47108
diff changeset
    32
section {* Operations on Predicates *}
35953
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)"
46884
154dc6ec0041 tuned proofs
noschinl
parents: 45970
diff changeset
    36
  by (simp add: fun_eq_iff)
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    37
36253
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36246
diff changeset
    38
lemma subset_eq[code_pred_inline]:
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36246
diff changeset
    39
  "(P :: 'a => bool) < (Q :: 'a => bool) == ((\<exists>x. Q x \<and> (\<not> P x)) \<and> (\<forall> x. P x --> Q x))"
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 45461
diff changeset
    40
  by (rule eq_reflection) (auto simp add: less_fun_def le_fun_def)
36253
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36246
diff changeset
    41
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    42
lemma set_equality[code_pred_inline]:
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 45461
diff changeset
    43
  "A = B \<longleftrightarrow> (\<forall>x. A x \<longrightarrow> B x) \<and> (\<forall>x. B x \<longrightarrow> A x)"
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 45461
diff changeset
    44
  by (auto simp add: fun_eq_iff)
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 45461
diff changeset
    45
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    46
section {* Setup for Numerals *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    47
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 51143
diff changeset
    48
setup {* Predicate_Compile_Data.ignore_consts [@{const_name numeral}] *}
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 51143
diff changeset
    49
setup {* Predicate_Compile_Data.keep_functions [@{const_name numeral}] *}
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    50
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    51
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
    52
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
    53
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
    54
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
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
    56
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 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
    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
  "plus_eq_nat x y z = (x + y = z)"
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    60
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
    61
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
    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
  "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
    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 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
    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
  "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
    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 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
    70
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
    71
  "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
    72
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
definition subtract
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    74
where
45231
d85a2fdc586c replacing code_inline by code_unfold, removing obsolete code_unfold, code_inline del now that the ancient code generator is removed
bulwahn
parents: 44928
diff changeset
    75
  [code_unfold]: "subtract x y = y - x"
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
    76
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
    77
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
    78
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
    79
  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
    80
  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
    81
  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
    82
  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
    83
  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
    84
  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
    85
  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
    86
  val ooi = Fun (Output, Fun (Output, Fun (Input, Bool)))
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39650
diff changeset
    87
  val plus_nat = Core_Data.functional_compilation @{const_name plus} iio
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39650
diff changeset
    88
  val minus_nat = Core_Data.functional_compilation @{const_name "minus"} iio
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
    89
  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
    90
    let
45461
130c90bb80b4 using more conventional names for monad plus operations
bulwahn
parents: 45231
diff changeset
    91
      val T = Predicate_Compile_Aux.mk_monadT compfuns @{typ nat}
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
    92
    in
44241
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 41075
diff changeset
    93
      absdummy @{typ nat} (absdummy @{typ nat}
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 41075
diff changeset
    94
        (Const (@{const_name "If"}, @{typ bool} --> T --> T --> T) $
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
    95
          (@{term "op > :: nat => nat => bool"} $ Bound 1 $ Bound 0) $
45461
130c90bb80b4 using more conventional names for monad plus operations
bulwahn
parents: 45231
diff changeset
    96
          Predicate_Compile_Aux.mk_empty compfuns @{typ nat} $
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
    97
          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
    98
          (@{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
    99
    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
   100
  fun enumerate_addups_nat compfuns (_ : typ) =
44241
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 41075
diff changeset
   101
    absdummy @{typ nat} (Predicate_Compile_Aux.mk_iterate_upto compfuns @{typ "nat * nat"}
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48640
diff changeset
   102
    (absdummy @{typ natural} (@{term "Pair :: nat => nat => nat * nat"} $
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48640
diff changeset
   103
      (@{term "nat_of_natural"} $ Bound 0) $
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48640
diff changeset
   104
      (@{term "op - :: nat => nat => nat"} $ Bound 1 $ (@{term "nat_of_natural"} $ Bound 0))),
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48640
diff changeset
   105
      @{term "0 :: natural"}, @{term "natural_of_nat"} $ Bound 0))
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
   106
  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
   107
    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
   108
      val (single_const, _) = strip_comb (Predicate_Compile_Aux.mk_single compfuns @{term "0 :: nat"})
45461
130c90bb80b4 using more conventional names for monad plus operations
bulwahn
parents: 45231
diff changeset
   109
      val T = Predicate_Compile_Aux.mk_monadT compfuns @{typ nat}
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
   110
    in
44241
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 41075
diff changeset
   111
      absdummy @{typ nat} (absdummy @{typ nat}
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 41075
diff changeset
   112
        (Const (@{const_name If}, @{typ bool} --> T --> T --> T) $
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
   113
          (@{term "op = :: nat => nat => bool"} $ Bound 0 $ @{term "0::nat"}) $
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48640
diff changeset
   114
          (Predicate_Compile_Aux.mk_iterate_upto compfuns @{typ nat} (@{term "nat_of_natural"},
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 48640
diff changeset
   115
            @{term "0::natural"}, @{term "natural_of_nat"} $ Bound 1)) $
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
   116
            (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
   117
    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
   118
in
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39650
diff changeset
   119
  Core_Data.force_modes_and_compilations @{const_name plus_eq_nat}
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
   120
    [(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
   121
     (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
   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 "plus :: nat => nat => nat"}, @{term "plus_eq_nat"})
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39650
diff changeset
   124
  #> Core_Data.force_modes_and_compilations @{const_name minus_eq_nat}
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
   125
       [(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
   126
  #> 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
   127
      (@{term "minus :: nat => nat => nat"}, @{term "minus_eq_nat"})
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39650
diff changeset
   128
  #> Core_Data.force_modes_and_functions @{const_name plus_eq_int}
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
   129
    [(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
   130
     (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
   131
  #> 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
   132
       (@{term "plus :: int => int => int"}, @{term "plus_eq_int"})
40054
cd7b1fa20bce adapting alternative_defs, predicate_compile_quickcheck, examples and code_prolog
bulwahn
parents: 39650
diff changeset
   133
  #> Core_Data.force_modes_and_functions @{const_name minus_eq_int}
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
   134
    [(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
   135
     (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
   136
  #> 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
   137
      (@{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
   138
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
   139
*}
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
   140
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
   141
subsection {* Inductive definitions for ordering on naturals *}
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   142
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   143
inductive less_nat
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   144
where
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   145
  "less_nat 0 (Suc y)"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   146
| "less_nat x y ==> less_nat (Suc x) (Suc y)"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   147
36246
43fecedff8cf added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
bulwahn
parents: 36053
diff changeset
   148
lemma less_nat[code_pred_inline]:
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   149
  "x < y = less_nat x y"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   150
apply (rule iffI)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   151
apply (induct x arbitrary: y)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   152
apply (case_tac y) apply (auto intro: less_nat.intros)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   153
apply (case_tac y)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   154
apply (auto intro: less_nat.intros)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   155
apply (induct rule: less_nat.induct)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   156
apply auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   157
done
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   158
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   159
inductive less_eq_nat
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   160
where
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   161
  "less_eq_nat 0 y"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   162
| "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
   163
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   164
lemma [code_pred_inline]:
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   165
"x <= y = less_eq_nat x y"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   166
apply (rule iffI)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   167
apply (induct x arbitrary: y)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   168
apply (auto intro: less_eq_nat.intros)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   169
apply (case_tac y) apply (auto intro: less_eq_nat.intros)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   170
apply (induct rule: less_eq_nat.induct)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   171
apply auto done
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   172
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   173
section {* Alternative list definitions *}
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   174
48640
053cc8dfde35 fixed document;
wenzelm
parents: 47840
diff changeset
   175
subsection {* Alternative rules for @{text length} *}
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
   176
56679
5545bfdfefcc avoid name shadowing
blanchet
parents: 56073
diff changeset
   177
definition size_list' :: "'a list => nat"
5545bfdfefcc avoid name shadowing
blanchet
parents: 56073
diff changeset
   178
where "size_list' = size"
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
   179
56679
5545bfdfefcc avoid name shadowing
blanchet
parents: 56073
diff changeset
   180
lemma size_list'_simps:
5545bfdfefcc avoid name shadowing
blanchet
parents: 56073
diff changeset
   181
  "size_list' [] = 0"
5545bfdfefcc avoid name shadowing
blanchet
parents: 56073
diff changeset
   182
  "size_list' (x # xs) = Suc (size_list' xs)"
5545bfdfefcc avoid name shadowing
blanchet
parents: 56073
diff changeset
   183
by (auto simp add: size_list'_def)
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
   184
56679
5545bfdfefcc avoid name shadowing
blanchet
parents: 56073
diff changeset
   185
declare size_list'_simps[code_pred_def]
5545bfdfefcc avoid name shadowing
blanchet
parents: 56073
diff changeset
   186
declare size_list'_def[symmetric, code_pred_inline]
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   187
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   188
48640
053cc8dfde35 fixed document;
wenzelm
parents: 47840
diff changeset
   189
subsection {* Alternative rules for @{text list_all2} *}
35953
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 list_all2_NilI [code_pred_intro]: "list_all2 P [] []"
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   192
by auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   193
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   194
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
   195
by auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   196
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   197
code_pred [skip_proof] list_all2
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   198
proof -
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   199
  case list_all2
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   200
  from this show thesis
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   201
    apply -
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   202
    apply (case_tac xb)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   203
    apply (case_tac xc)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   204
    apply auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   205
    apply (case_tac xc)
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   206
    apply auto
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   207
    done
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   208
qed
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   209
40548
54eb5fd36e5e ignoring the constant STR in the predicate compiler
bulwahn
parents: 40054
diff changeset
   210
section {* Setup for String.literal *}
54eb5fd36e5e ignoring the constant STR in the predicate compiler
bulwahn
parents: 40054
diff changeset
   211
54eb5fd36e5e ignoring the constant STR in the predicate compiler
bulwahn
parents: 40054
diff changeset
   212
setup {* Predicate_Compile_Data.ignore_consts [@{const_name "STR"}] *}
54eb5fd36e5e ignoring the constant STR in the predicate compiler
bulwahn
parents: 40054
diff changeset
   213
36246
43fecedff8cf added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
bulwahn
parents: 36053
diff changeset
   214
section {* Simplification rules for optimisation *}
43fecedff8cf added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
bulwahn
parents: 36053
diff changeset
   215
43fecedff8cf added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
bulwahn
parents: 36053
diff changeset
   216
lemma [code_pred_simp]: "\<not> False == True"
43fecedff8cf added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
bulwahn
parents: 36053
diff changeset
   217
by auto
43fecedff8cf added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
bulwahn
parents: 36053
diff changeset
   218
43fecedff8cf added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
bulwahn
parents: 36053
diff changeset
   219
lemma [code_pred_simp]: "\<not> True == False"
43fecedff8cf added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
bulwahn
parents: 36053
diff changeset
   220
by auto
43fecedff8cf added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
bulwahn
parents: 36053
diff changeset
   221
43fecedff8cf added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
bulwahn
parents: 36053
diff changeset
   222
lemma less_nat_k_0 [code_pred_simp]: "less_nat k 0 == False"
43fecedff8cf added peephole optimisations to the predicate compiler; added structure Predicate_Compile_Simps for peephole optimisations
bulwahn
parents: 36053
diff changeset
   223
unfolding less_nat[symmetric] by auto
35953
0460ff79bb52 moved further predicate compile files to HOL-Library
bulwahn
parents:
diff changeset
   224
46884
154dc6ec0041 tuned proofs
noschinl
parents: 45970
diff changeset
   225
end