src/HOL/Predicate_Compile_Examples/Predicate_Compile_Examples.thy
author bulwahn
Mon, 20 Sep 2010 09:26:20 +0200
changeset 39545 631cf48c7894
parent 39543 9ff9651757cd
child 39648 655307cb8489
permissions -rw-r--r--
code_pred_intro can be used to name facts for the code_pred command
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
     1
theory Predicate_Compile_Examples
35954
d87d85a5d9ab adopting examples to Library move
bulwahn
parents: 35950
diff changeset
     2
imports Predicate_Compile_Alternative_Defs
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
     3
begin
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
     4
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
     5
subsection {* Basic predicates *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
     6
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
     7
inductive False' :: "bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
     8
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
     9
code_pred (expected_modes: bool) False' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    10
code_pred [dseq] False' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    11
code_pred [random_dseq] False' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    12
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    13
values [expected "{}" pred] "{x. False'}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    14
values [expected "{}" dseq 1] "{x. False'}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    15
values [expected "{}" random_dseq 1, 1, 1] "{x. False'}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    16
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    17
value "False'"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    18
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    19
inductive True' :: "bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    20
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    21
  "True ==> True'"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    22
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    23
code_pred True' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    24
code_pred [dseq] True' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    25
code_pred [random_dseq] True' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    26
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    27
thm True'.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    28
thm True'.dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    29
thm True'.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    30
values [expected "{()}" ]"{x. True'}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    31
values [expected "{}" dseq 0] "{x. True'}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    32
values [expected "{()}" dseq 1] "{x. True'}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    33
values [expected "{()}" dseq 2] "{x. True'}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    34
values [expected "{}" random_dseq 1, 1, 0] "{x. True'}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    35
values [expected "{}" random_dseq 1, 1, 1] "{x. True'}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    36
values [expected "{()}" random_dseq 1, 1, 2] "{x. True'}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    37
values [expected "{()}" random_dseq 1, 1, 3] "{x. True'}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    38
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    39
inductive EmptySet :: "'a \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    40
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    41
code_pred (expected_modes: o => bool, i => bool) EmptySet .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    42
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    43
definition EmptySet' :: "'a \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    44
where "EmptySet' = {}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    45
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    46
code_pred (expected_modes: o => bool, i => bool) [inductify] EmptySet' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    47
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    48
inductive EmptyRel :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    49
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    50
code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) EmptyRel .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    51
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    52
inductive EmptyClosure :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    53
for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    54
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    55
code_pred
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    56
  (expected_modes: (o => o => bool) => o => o => bool, (o => o => bool) => i => o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    57
         (o => o => bool) => o => i => bool, (o => o => bool) => i => i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    58
         (i => o => bool) => o => o => bool, (i => o => bool) => i => o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    59
         (i => o => bool) => o => i => bool, (i => o => bool) => i => i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    60
         (o => i => bool) => o => o => bool, (o => i => bool) => i => o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    61
         (o => i => bool) => o => i => bool, (o => i => bool) => i => i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    62
         (i => i => bool) => o => o => bool, (i => i => bool) => i => o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    63
         (i => i => bool) => o => i => bool, (i => i => bool) => i => i => bool)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    64
  EmptyClosure .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    65
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    66
thm EmptyClosure.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    67
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    68
(* TODO: inductive package is broken!
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    69
inductive False'' :: "bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    70
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    71
  "False \<Longrightarrow> False''"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    72
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    73
code_pred (expected_modes: []) False'' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    74
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    75
inductive EmptySet'' :: "'a \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    76
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    77
  "False \<Longrightarrow> EmptySet'' x"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    78
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    79
code_pred (expected_modes: [1]) EmptySet'' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    80
code_pred (expected_modes: [], [1]) [inductify] EmptySet'' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    81
*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    82
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    83
consts a' :: 'a
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    84
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    85
inductive Fact :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    86
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    87
"Fact a' a'"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    88
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    89
code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) Fact .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    90
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    91
inductive zerozero :: "nat * nat => bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    92
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    93
  "zerozero (0, 0)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    94
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    95
code_pred (expected_modes: i => bool, i * o => bool, o * i => bool, o => bool) zerozero .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    96
code_pred [dseq] zerozero .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    97
code_pred [random_dseq] zerozero .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    98
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
    99
thm zerozero.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   100
thm zerozero.dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   101
thm zerozero.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   102
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   103
text {* We expect the user to expand the tuples in the values command.
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   104
The following values command is not supported. *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   105
(*values "{x. zerozero x}" *)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   106
text {* Instead, the user must type *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   107
values "{(x, y). zerozero (x, y)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   108
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   109
values [expected "{}" dseq 0] "{(x, y). zerozero (x, y)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   110
values [expected "{(0::nat, 0::nat)}" dseq 1] "{(x, y). zerozero (x, y)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   111
values [expected "{(0::nat, 0::nat)}" dseq 2] "{(x, y). zerozero (x, y)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   112
values [expected "{}" random_dseq 1, 1, 2] "{(x, y). zerozero (x, y)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   113
values [expected "{(0::nat, 0:: nat)}" random_dseq 1, 1, 3] "{(x, y). zerozero (x, y)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   114
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   115
inductive nested_tuples :: "((int * int) * int * int) => bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   116
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   117
  "nested_tuples ((0, 1), 2, 3)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   118
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   119
code_pred nested_tuples .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   120
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   121
inductive JamesBond :: "nat => int => code_numeral => bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   122
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   123
  "JamesBond 0 0 7"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   124
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   125
code_pred JamesBond .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   126
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   127
values [expected "{(0::nat, 0::int , 7::code_numeral)}"] "{(a, b, c). JamesBond a b c}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   128
values [expected "{(0::nat, 7::code_numeral, 0:: int)}"] "{(a, c, b). JamesBond a b c}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   129
values [expected "{(0::int, 0::nat, 7::code_numeral)}"] "{(b, a, c). JamesBond a b c}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   130
values [expected "{(0::int, 7::code_numeral, 0::nat)}"] "{(b, c, a). JamesBond a b c}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   131
values [expected "{(7::code_numeral, 0::nat, 0::int)}"] "{(c, a, b). JamesBond a b c}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   132
values [expected "{(7::code_numeral, 0::int, 0::nat)}"] "{(c, b, a). JamesBond a b c}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   133
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   134
values [expected "{(7::code_numeral, 0::int)}"] "{(a, b). JamesBond 0 b a}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   135
values [expected "{(7::code_numeral, 0::nat)}"] "{(c, a). JamesBond a 0 c}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   136
values [expected "{(0::nat, 7::code_numeral)}"] "{(a, c). JamesBond a 0 c}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   137
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   138
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   139
subsection {* Alternative Rules *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   140
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   141
datatype char = C | D | E | F | G | H
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   142
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   143
inductive is_C_or_D
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   144
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   145
  "(x = C) \<or> (x = D) ==> is_C_or_D x"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   146
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   147
code_pred (expected_modes: i => bool) is_C_or_D .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   148
thm is_C_or_D.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   149
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   150
inductive is_D_or_E
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   151
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   152
  "(x = D) \<or> (x = E) ==> is_D_or_E x"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   153
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   154
lemma [code_pred_intro]:
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   155
  "is_D_or_E D"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   156
by (auto intro: is_D_or_E.intros)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   157
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   158
lemma [code_pred_intro]:
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   159
  "is_D_or_E E"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   160
by (auto intro: is_D_or_E.intros)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   161
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   162
code_pred (expected_modes: o => bool, i => bool) is_D_or_E
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   163
proof -
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   164
  case is_D_or_E
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   165
  from this(1) show thesis
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   166
  proof
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   167
    fix xa
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   168
    assume x: "x = xa"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   169
    assume "xa = D \<or> xa = E"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   170
    from this show thesis
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   171
    proof
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   172
      assume "xa = D" from this x is_D_or_E(2) show thesis by simp
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   173
    next
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   174
      assume "xa = E" from this x is_D_or_E(3) show thesis by simp
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   175
    qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   176
  qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   177
qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   178
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   179
thm is_D_or_E.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   180
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   181
inductive is_F_or_G
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   182
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   183
  "x = F \<or> x = G ==> is_F_or_G x"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   184
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   185
lemma [code_pred_intro]:
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   186
  "is_F_or_G F"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   187
by (auto intro: is_F_or_G.intros)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   188
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   189
lemma [code_pred_intro]:
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   190
  "is_F_or_G G"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   191
by (auto intro: is_F_or_G.intros)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   192
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   193
inductive is_FGH
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   194
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   195
  "is_F_or_G x ==> is_FGH x"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   196
| "is_FGH H"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   197
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   198
text {* Compilation of is_FGH requires elimination rule for is_F_or_G *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   199
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   200
code_pred (expected_modes: o => bool, i => bool) is_FGH
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   201
proof -
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   202
  case is_F_or_G
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   203
  from this(1) show thesis
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   204
  proof
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   205
    fix xa
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   206
    assume x: "x = xa"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   207
    assume "xa = F \<or> xa = G"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   208
    from this show thesis
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   209
    proof
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   210
      assume "xa = F"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   211
      from this x is_F_or_G(2) show thesis by simp
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   212
    next
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   213
      assume "xa = G"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   214
      from this x is_F_or_G(3) show thesis by simp
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   215
    qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   216
  qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   217
qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   218
39545
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   219
subsection {* Named alternative rules *}
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   220
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   221
inductive appending
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   222
where
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   223
  nil: "appending [] ys ys"
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   224
| cons: "appending xs ys zs \<Longrightarrow> appending (x#xs) ys (x#zs)"
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   225
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   226
lemma appending_alt_nil: "ys = zs \<Longrightarrow> appending [] ys zs"
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   227
by (auto intro: appending.intros)
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   228
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   229
lemma appending_alt_cons: "xs' = x # xs \<Longrightarrow> appending xs ys zs \<Longrightarrow> zs' = x # zs \<Longrightarrow> appending xs' ys zs'"
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   230
by (auto intro: appending.intros)
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   231
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   232
text {* With code_pred_intro, we can give fact names to the alternative rules,
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   233
  which are used for the code_pred command. *}
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   234
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   235
declare appending_alt_nil[code_pred_intro alt_nil] appending_alt_cons[code_pred_intro alt_cons]
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   236
 
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   237
code_pred appending
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   238
proof -
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   239
  case appending
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   240
  from prems show thesis
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   241
  proof(cases)
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   242
    case nil
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   243
    from alt_nil nil show thesis by auto
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   244
  next
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   245
    case cons
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   246
    from alt_cons cons show thesis by fastsimp
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   247
  qed
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   248
qed
631cf48c7894 code_pred_intro can be used to name facts for the code_pred command
bulwahn
parents: 39543
diff changeset
   249
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   250
subsection {* Preprocessor Inlining  *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   251
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   252
definition "equals == (op =)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   253
 
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   254
inductive zerozero' :: "nat * nat => bool" where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   255
  "equals (x, y) (0, 0) ==> zerozero' (x, y)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   256
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   257
code_pred (expected_modes: i => bool) zerozero' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   258
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   259
lemma zerozero'_eq: "zerozero' x == zerozero x"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   260
proof -
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   261
  have "zerozero' = zerozero"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   262
    apply (auto simp add: mem_def)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   263
    apply (cases rule: zerozero'.cases)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   264
    apply (auto simp add: equals_def intro: zerozero.intros)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   265
    apply (cases rule: zerozero.cases)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   266
    apply (auto simp add: equals_def intro: zerozero'.intros)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   267
    done
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   268
  from this show "zerozero' x == zerozero x" by auto
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   269
qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   270
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   271
declare zerozero'_eq [code_pred_inline]
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   272
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   273
definition "zerozero'' x == zerozero' x"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   274
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   275
text {* if preprocessing fails, zerozero'' will not have all modes. *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   276
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   277
code_pred (expected_modes: i * i => bool, i * o => bool, o * i => bool, o => bool) [inductify] zerozero'' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   278
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   279
subsection {* Sets and Numerals *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   280
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   281
definition
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   282
  "one_or_two = {Suc 0, (Suc (Suc 0))}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   283
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   284
code_pred [inductify] one_or_two .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   285
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   286
code_pred [dseq] one_or_two .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   287
code_pred [random_dseq] one_or_two .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   288
thm one_or_two.dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   289
values [expected "{Suc 0::nat, 2::nat}"] "{x. one_or_two x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   290
values [random_dseq 0,0,10] 3 "{x. one_or_two x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   291
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   292
inductive one_or_two' :: "nat => bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   293
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   294
  "one_or_two' 1"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   295
| "one_or_two' 2"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   296
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   297
code_pred one_or_two' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   298
thm one_or_two'.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   299
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   300
values "{x. one_or_two' x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   301
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   302
definition one_or_two'':
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   303
  "one_or_two'' == {1, (2::nat)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   304
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   305
code_pred [inductify] one_or_two'' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   306
thm one_or_two''.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   307
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   308
values "{x. one_or_two'' x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   309
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   310
subsection {* even predicate *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   311
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   312
inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   313
    "even 0"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   314
  | "even n \<Longrightarrow> odd (Suc n)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   315
  | "odd n \<Longrightarrow> even (Suc n)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   316
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   317
code_pred (expected_modes: i => bool, o => bool) even .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   318
code_pred [dseq] even .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   319
code_pred [random_dseq] even .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   320
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   321
thm odd.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   322
thm even.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   323
thm odd.dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   324
thm even.dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   325
thm odd.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   326
thm even.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   327
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   328
values "{x. even 2}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   329
values "{x. odd 2}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   330
values 10 "{n. even n}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   331
values 10 "{n. odd n}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   332
values [expected "{}" dseq 2] "{x. even 6}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   333
values [expected "{}" dseq 6] "{x. even 6}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   334
values [expected "{()}" dseq 7] "{x. even 6}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   335
values [dseq 2] "{x. odd 7}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   336
values [dseq 6] "{x. odd 7}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   337
values [dseq 7] "{x. odd 7}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   338
values [expected "{()}" dseq 8] "{x. odd 7}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   339
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   340
values [expected "{}" dseq 0] 8 "{x. even x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   341
values [expected "{0::nat}" dseq 1] 8 "{x. even x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   342
values [expected "{0::nat, 2}" dseq 3] 8 "{x. even x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   343
values [expected "{0::nat, 2}" dseq 4] 8 "{x. even x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   344
values [expected "{0::nat, 2, 4}" dseq 6] 8 "{x. even x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   345
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   346
values [random_dseq 1, 1, 0] 8 "{x. even x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   347
values [random_dseq 1, 1, 1] 8 "{x. even x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   348
values [random_dseq 1, 1, 2] 8 "{x. even x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   349
values [random_dseq 1, 1, 3] 8 "{x. even x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   350
values [random_dseq 1, 1, 6] 8 "{x. even x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   351
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   352
values [expected "{}" random_dseq 1, 1, 7] "{x. odd 7}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   353
values [random_dseq 1, 1, 8] "{x. odd 7}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   354
values [random_dseq 1, 1, 9] "{x. odd 7}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   355
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   356
definition odd' where "odd' x == \<not> even x"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   357
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   358
code_pred (expected_modes: i => bool) [inductify] odd' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   359
code_pred [dseq inductify] odd' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   360
code_pred [random_dseq inductify] odd' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   361
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   362
values [expected "{}" dseq 2] "{x. odd' 7}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   363
values [expected "{()}" dseq 9] "{x. odd' 7}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   364
values [expected "{}" dseq 2] "{x. odd' 8}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   365
values [expected "{}" dseq 10] "{x. odd' 8}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   366
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   367
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   368
inductive is_even :: "nat \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   369
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   370
  "n mod 2 = 0 \<Longrightarrow> is_even n"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   371
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   372
code_pred (expected_modes: i => bool) is_even .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   373
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   374
subsection {* append predicate *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   375
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   376
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   377
    "append [] xs xs"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   378
  | "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   379
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   380
code_pred (modes: i => i => o => bool as "concat", o => o => i => bool as "slice", o => i => i => bool as prefix,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   381
  i => o => i => bool as suffix, i => i => i => bool) append .
38664
7215ae18f44b added support for xsymbol syntax for mode annotations in code_pred command
bulwahn
parents: 36511
diff changeset
   382
code_pred (modes: i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool as "concat", o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as "slice", o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool as prefix,
7215ae18f44b added support for xsymbol syntax for mode annotations in code_pred command
bulwahn
parents: 36511
diff changeset
   383
  i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as suffix, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool) append .
7215ae18f44b added support for xsymbol syntax for mode annotations in code_pred command
bulwahn
parents: 36511
diff changeset
   384
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   385
code_pred [dseq] append .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   386
code_pred [random_dseq] append .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   387
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   388
thm append.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   389
thm append.dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   390
thm append.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   391
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   392
values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   393
values "{zs. append [0, Suc 0, 2] [17, 8] zs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   394
values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0, 5]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   395
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   396
values [expected "{}" dseq 0] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   397
values [expected "{(([]::nat list), [Suc 0, 2, 3, 4, (5::nat)])}" dseq 1] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   398
values [dseq 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   399
values [dseq 6] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   400
values [random_dseq 1, 1, 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   401
values [random_dseq 1, 1, 1] 10 "{(xs, ys, zs::int list). append xs ys zs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   402
values [random_dseq 1, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   403
values [random_dseq 3, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   404
values [random_dseq 1, 3, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   405
values [random_dseq 1, 1, 4] 10 "{(xs, ys, zs::int list). append xs ys zs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   406
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   407
value [code] "Predicate.the (concat [0::int, 1, 2] [3, 4, 5])"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   408
value [code] "Predicate.the (slice ([]::int list))"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   409
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   410
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   411
text {* tricky case with alternative rules *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   412
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   413
inductive append2
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   414
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   415
  "append2 [] xs xs"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   416
| "append2 xs ys zs \<Longrightarrow> append2 (x # xs) ys (x # zs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   417
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   418
lemma append2_Nil: "append2 [] (xs::'b list) xs"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   419
  by (simp add: append2.intros(1))
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   420
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   421
lemmas [code_pred_intro] = append2_Nil append2.intros(2)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   422
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   423
code_pred (expected_modes: i => i => o => bool, o => o => i => bool, o => i => i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   424
  i => o => i => bool, i => i => i => bool) append2
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   425
proof -
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   426
  case append2
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   427
  from append2(1) show thesis
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   428
  proof
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   429
    fix xs
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   430
    assume "xa = []" "xb = xs" "xc = xs"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   431
    from this append2(2) show thesis by simp
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   432
  next
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   433
    fix xs ys zs x
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   434
    assume "xa = x # xs" "xb = ys" "xc = x # zs" "append2 xs ys zs"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   435
    from this append2(3) show thesis by fastsimp
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   436
  qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   437
qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   438
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   439
inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   440
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   441
  "tupled_append ([], xs, xs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   442
| "tupled_append (xs, ys, zs) \<Longrightarrow> tupled_append (x # xs, ys, x # zs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   443
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   444
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   445
  i * o * i => bool, i * i * i => bool) tupled_append .
38664
7215ae18f44b added support for xsymbol syntax for mode annotations in code_pred command
bulwahn
parents: 36511
diff changeset
   446
7215ae18f44b added support for xsymbol syntax for mode annotations in code_pred command
bulwahn
parents: 36511
diff changeset
   447
code_pred (expected_modes: i \<times> i \<times> o \<Rightarrow> bool, o \<times> o \<times> i \<Rightarrow> bool, o \<times> i \<times> i \<Rightarrow> bool,
7215ae18f44b added support for xsymbol syntax for mode annotations in code_pred command
bulwahn
parents: 36511
diff changeset
   448
  i \<times> o \<times> i \<Rightarrow> bool, i \<times> i \<times> i \<Rightarrow> bool) tupled_append .
7215ae18f44b added support for xsymbol syntax for mode annotations in code_pred command
bulwahn
parents: 36511
diff changeset
   449
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   450
code_pred [random_dseq] tupled_append .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   451
thm tupled_append.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   452
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   453
values "{xs. tupled_append ([(1::nat), 2, 3], [4, 5], xs)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   454
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   455
inductive tupled_append'
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   456
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   457
"tupled_append' ([], xs, xs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   458
| "[| ys = fst (xa, y); x # zs = snd (xa, y);
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   459
 tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   460
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   461
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   462
  i * o * i => bool, i * i * i => bool) tupled_append' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   463
thm tupled_append'.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   464
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   465
inductive tupled_append'' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   466
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   467
  "tupled_append'' ([], xs, xs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   468
| "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, zs) \<Longrightarrow> tupled_append'' (x # xs, yszs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   469
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   470
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   471
  i * o * i => bool, i * i * i => bool) tupled_append'' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   472
thm tupled_append''.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   473
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   474
inductive tupled_append''' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   475
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   476
  "tupled_append''' ([], xs, xs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   477
| "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \<Longrightarrow> tupled_append''' (x # xs, ys, x # zs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   478
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   479
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   480
  i * o * i => bool, i * i * i => bool) tupled_append''' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   481
thm tupled_append'''.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   482
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   483
subsection {* map_ofP predicate *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   484
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   485
inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   486
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   487
  "map_ofP ((a, b)#xs) a b"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   488
| "map_ofP xs a b \<Longrightarrow> map_ofP (x#xs) a b"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   489
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   490
code_pred (expected_modes: i => o => o => bool, i => i => o => bool, i => o => i => bool, i => i => i => bool) map_ofP .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   491
thm map_ofP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   492
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   493
subsection {* filter predicate *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   494
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   495
inductive filter1
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   496
for P
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   497
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   498
  "filter1 P [] []"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   499
| "P x ==> filter1 P xs ys ==> filter1 P (x#xs) (x#ys)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   500
| "\<not> P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   501
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   502
code_pred (expected_modes: (i => bool) => i => o => bool, (i => bool) => i => i => bool) filter1 .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   503
code_pred [dseq] filter1 .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   504
code_pred [random_dseq] filter1 .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   505
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   506
thm filter1.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   507
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   508
values [expected "{[0::nat, 2, 4]}"] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   509
values [expected "{}" dseq 9] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   510
values [expected "{[0::nat, 2, 4]}" dseq 10] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   511
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   512
inductive filter2
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   513
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   514
  "filter2 P [] []"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   515
| "P x ==> filter2 P xs ys ==> filter2 P (x#xs) (x#ys)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   516
| "\<not> P x ==> filter2 P xs ys ==> filter2 P (x#xs) ys"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   517
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   518
code_pred (expected_modes: (i => bool) => i => i => bool, (i => bool) => i => o => bool) filter2 .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   519
code_pred [dseq] filter2 .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   520
code_pred [random_dseq] filter2 .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   521
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   522
thm filter2.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   523
thm filter2.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   524
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   525
(*
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   526
inductive filter3
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   527
for P
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   528
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   529
  "List.filter P xs = ys ==> filter3 P xs ys"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   530
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   531
code_pred (expected_modes: (o => bool) => i => o => bool, (o => bool) => i => i => bool , (i => bool) => i => o => bool, (i => bool) => i => i => bool) [skip_proof] filter3 .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   532
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   533
code_pred [dseq] filter3 .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   534
thm filter3.dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   535
*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   536
(*
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   537
inductive filter4
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   538
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   539
  "List.filter P xs = ys ==> filter4 P xs ys"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   540
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   541
code_pred (expected_modes: i => i => o => bool, i => i => i => bool) filter4 .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   542
(*code_pred [depth_limited] filter4 .*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   543
(*code_pred [random] filter4 .*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   544
*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   545
subsection {* reverse predicate *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   546
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   547
inductive rev where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   548
    "rev [] []"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   549
  | "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   550
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   551
code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) rev .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   552
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   553
thm rev.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   554
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   555
values "{xs. rev [0, 1, 2, 3::nat] xs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   556
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   557
inductive tupled_rev where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   558
  "tupled_rev ([], [])"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   559
| "tupled_rev (xs, xs') \<Longrightarrow> tupled_append (xs', [x], ys) \<Longrightarrow> tupled_rev (x#xs, ys)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   560
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   561
code_pred (expected_modes: i * o => bool, o * i => bool, i * i => bool) tupled_rev .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   562
thm tupled_rev.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   563
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   564
subsection {* partition predicate *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   565
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   566
inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   567
  for f where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   568
    "partition f [] [] []"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   569
  | "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   570
  | "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   571
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   572
code_pred (expected_modes: (i => bool) => i => o => o => bool, (i => bool) => o => i => i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   573
  (i => bool) => i => i => o => bool, (i => bool) => i => o => i => bool, (i => bool) => i => i => i => bool)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   574
  partition .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   575
code_pred [dseq] partition .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   576
code_pred [random_dseq] partition .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   577
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   578
values 10 "{(ys, zs). partition is_even
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   579
  [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   580
values 10 "{zs. partition is_even zs [0, 2] [3, 5]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   581
values 10 "{zs. partition is_even zs [0, 7] [3, 5]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   582
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   583
inductive tupled_partition :: "('a \<Rightarrow> bool) \<Rightarrow> ('a list \<times> 'a list \<times> 'a list) \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   584
  for f where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   585
   "tupled_partition f ([], [], [])"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   586
  | "f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, x # ys, zs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   587
  | "\<not> f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, ys, x # zs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   588
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   589
code_pred (expected_modes: (i => bool) => i => bool, (i => bool) => (i * i * o) => bool, (i => bool) => (i * o * i) => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   590
  (i => bool) => (o * i * i) => bool, (i => bool) => (i * o * o) => bool) tupled_partition .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   591
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   592
thm tupled_partition.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   593
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   594
lemma [code_pred_intro]:
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   595
  "r a b \<Longrightarrow> tranclp r a b"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   596
  "r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   597
  by auto
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   598
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   599
subsection {* transitive predicate *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   600
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   601
text {* Also look at the tabled transitive closure in the Library *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   602
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   603
code_pred (modes: (i => o => bool) => i => i => bool, (i => o => bool) => i => o => bool as forwards_trancl,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   604
  (o => i => bool) => i => i => bool, (o => i => bool) => o => i => bool as backwards_trancl, (o => o => bool) => i => i => bool, (o => o => bool) => i => o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   605
  (o => o => bool) => o => i => bool, (o => o => bool) => o => o => bool) tranclp
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   606
proof -
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   607
  case tranclp
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   608
  from this converse_tranclpE[OF this(1)] show thesis by metis
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   609
qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   610
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   611
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   612
code_pred [dseq] tranclp .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   613
code_pred [random_dseq] tranclp .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   614
thm tranclp.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   615
thm tranclp.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   616
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   617
inductive rtrancl' :: "'a => 'a => ('a => 'a => bool) => bool" 
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   618
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   619
  "rtrancl' x x r"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   620
| "r x y ==> rtrancl' y z r ==> rtrancl' x z r"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   621
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   622
code_pred [random_dseq] rtrancl' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   623
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   624
thm rtrancl'.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   625
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   626
inductive rtrancl'' :: "('a * 'a * ('a \<Rightarrow> 'a \<Rightarrow> bool)) \<Rightarrow> bool"  
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   627
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   628
  "rtrancl'' (x, x, r)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   629
| "r x y \<Longrightarrow> rtrancl'' (y, z, r) \<Longrightarrow> rtrancl'' (x, z, r)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   630
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   631
code_pred rtrancl'' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   632
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   633
inductive rtrancl''' :: "('a * ('a * 'a) * ('a * 'a => bool)) => bool" 
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   634
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   635
  "rtrancl''' (x, (x, x), r)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   636
| "r (x, y) ==> rtrancl''' (y, (z, z), r) ==> rtrancl''' (x, (z, z), r)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   637
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   638
code_pred rtrancl''' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   639
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   640
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   641
inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   642
    "succ 0 1"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   643
  | "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   644
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   645
code_pred (modes: i => i => bool, i => o => bool, o => i => bool, o => o => bool) succ .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   646
code_pred [random_dseq] succ .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   647
thm succ.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   648
thm succ.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   649
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   650
values 10 "{(m, n). succ n m}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   651
values "{m. succ 0 m}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   652
values "{m. succ m 0}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   653
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   654
text {* values command needs mode annotation of the parameter succ
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   655
to disambiguate which mode is to be chosen. *} 
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   656
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   657
values [mode: i => o => bool] 20 "{n. tranclp succ 10 n}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   658
values [mode: o => i => bool] 10 "{n. tranclp succ n 10}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   659
values 20 "{(n, m). tranclp succ n m}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   660
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   661
inductive example_graph :: "int => int => bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   662
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   663
  "example_graph 0 1"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   664
| "example_graph 1 2"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   665
| "example_graph 1 3"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   666
| "example_graph 4 7"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   667
| "example_graph 4 5"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   668
| "example_graph 5 6"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   669
| "example_graph 7 6"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   670
| "example_graph 7 8"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   671
 
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   672
inductive not_reachable_in_example_graph :: "int => int => bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   673
where "\<not> (tranclp example_graph x y) ==> not_reachable_in_example_graph x y"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   674
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   675
code_pred (expected_modes: i => i => bool) not_reachable_in_example_graph .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   676
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   677
thm not_reachable_in_example_graph.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   678
thm tranclp.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   679
value "not_reachable_in_example_graph 0 3"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   680
value "not_reachable_in_example_graph 4 8"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   681
value "not_reachable_in_example_graph 5 6"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   682
text {* rtrancl compilation is strange! *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   683
(*
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   684
value "not_reachable_in_example_graph 0 4"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   685
value "not_reachable_in_example_graph 1 6"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   686
value "not_reachable_in_example_graph 8 4"*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   687
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   688
code_pred [dseq] not_reachable_in_example_graph .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   689
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   690
values [dseq 6] "{x. tranclp example_graph 0 3}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   691
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   692
values [dseq 0] "{x. not_reachable_in_example_graph 0 3}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   693
values [dseq 0] "{x. not_reachable_in_example_graph 0 4}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   694
values [dseq 20] "{x. not_reachable_in_example_graph 0 4}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   695
values [dseq 6] "{x. not_reachable_in_example_graph 0 3}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   696
values [dseq 3] "{x. not_reachable_in_example_graph 4 2}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   697
values [dseq 6] "{x. not_reachable_in_example_graph 4 2}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   698
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   699
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   700
inductive not_reachable_in_example_graph' :: "int => int => bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   701
where "\<not> (rtranclp example_graph x y) ==> not_reachable_in_example_graph' x y"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   702
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   703
code_pred not_reachable_in_example_graph' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   704
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   705
value "not_reachable_in_example_graph' 0 3"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   706
(* value "not_reachable_in_example_graph' 0 5" would not terminate *)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   707
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   708
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   709
(*values [depth_limited 0] "{x. not_reachable_in_example_graph' 0 3}"*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   710
(*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   711
(*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   712
(*values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   713
(*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   714
(*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   715
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   716
code_pred [dseq] not_reachable_in_example_graph' .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   717
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   718
(*thm not_reachable_in_example_graph'.dseq_equation*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   719
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   720
(*values [dseq 0] "{x. not_reachable_in_example_graph' 0 3}"*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   721
(*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   722
(*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   723
values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   724
(*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   725
(*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   726
36511
db2953f5887a added an example with a free function variable to the Predicate Compile examples
bulwahn
parents: 36260
diff changeset
   727
subsection {* Free function variable *}
db2953f5887a added an example with a free function variable to the Predicate Compile examples
bulwahn
parents: 36260
diff changeset
   728
db2953f5887a added an example with a free function variable to the Predicate Compile examples
bulwahn
parents: 36260
diff changeset
   729
inductive FF :: "nat => nat => bool"
db2953f5887a added an example with a free function variable to the Predicate Compile examples
bulwahn
parents: 36260
diff changeset
   730
where
db2953f5887a added an example with a free function variable to the Predicate Compile examples
bulwahn
parents: 36260
diff changeset
   731
  "f x = y ==> FF x y"
db2953f5887a added an example with a free function variable to the Predicate Compile examples
bulwahn
parents: 36260
diff changeset
   732
db2953f5887a added an example with a free function variable to the Predicate Compile examples
bulwahn
parents: 36260
diff changeset
   733
code_pred FF .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   734
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   735
subsection {* IMP *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   736
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   737
types
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   738
  var = nat
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   739
  state = "int list"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   740
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   741
datatype com =
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   742
  Skip |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   743
  Ass var "state => int" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   744
  Seq com com |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   745
  IF "state => bool" com com |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   746
  While "state => bool" com
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   747
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   748
inductive exec :: "com => state => state => bool" where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   749
"exec Skip s s" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   750
"exec (Ass x e) s (s[x := e(s)])" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   751
"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   752
"b s ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   753
"~b s ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   754
"~b s ==> exec (While b c) s s" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   755
"b s1 ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   756
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   757
code_pred exec .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   758
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   759
values "{t. exec
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   760
 (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))))
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   761
 [3,5] t}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   762
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   763
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   764
inductive tupled_exec :: "(com \<times> state \<times> state) \<Rightarrow> bool" where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   765
"tupled_exec (Skip, s, s)" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   766
"tupled_exec (Ass x e, s, s[x := e(s)])" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   767
"tupled_exec (c1, s1, s2) ==> tupled_exec (c2, s2, s3) ==> tupled_exec (Seq c1 c2, s1, s3)" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   768
"b s ==> tupled_exec (c1, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   769
"~b s ==> tupled_exec (c2, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   770
"~b s ==> tupled_exec (While b c, s, s)" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   771
"b s1 ==> tupled_exec (c, s1, s2) ==> tupled_exec (While b c, s2, s3) ==> tupled_exec (While b c, s1, s3)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   772
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   773
code_pred tupled_exec .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   774
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   775
values "{s. tupled_exec (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))), [3, 5], s)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   776
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   777
subsection {* CCS *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   778
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   779
text{* This example formalizes finite CCS processes without communication or
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   780
recursion. For simplicity, labels are natural numbers. *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   781
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   782
datatype proc = nil | pre nat proc | or proc proc | par proc proc
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   783
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   784
inductive step :: "proc \<Rightarrow> nat \<Rightarrow> proc \<Rightarrow> bool" where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   785
"step (pre n p) n p" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   786
"step p1 a q \<Longrightarrow> step (or p1 p2) a q" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   787
"step p2 a q \<Longrightarrow> step (or p1 p2) a q" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   788
"step p1 a q \<Longrightarrow> step (par p1 p2) a (par q p2)" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   789
"step p2 a q \<Longrightarrow> step (par p1 p2) a (par p1 q)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   790
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   791
code_pred step .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   792
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   793
inductive steps where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   794
"steps p [] p" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   795
"step p a q \<Longrightarrow> steps q as r \<Longrightarrow> steps p (a#as) r"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   796
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   797
code_pred steps .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   798
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   799
values 3 
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   800
 "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   801
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   802
values 5
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   803
 "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   804
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   805
values 3 "{(a,q). step (par nil nil) a q}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   806
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   807
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   808
inductive tupled_step :: "(proc \<times> nat \<times> proc) \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   809
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   810
"tupled_step (pre n p, n, p)" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   811
"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   812
"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   813
"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par q p2)" |
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   814
"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par p1 q)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   815
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   816
code_pred tupled_step .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   817
thm tupled_step.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   818
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   819
subsection {* divmod *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   820
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   821
inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   822
    "k < l \<Longrightarrow> divmod_rel k l 0 k"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   823
  | "k \<ge> l \<Longrightarrow> divmod_rel (k - l) l q r \<Longrightarrow> divmod_rel k l (Suc q) r"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   824
39438
c5ece2a7a86e Isar "default" step needs to fail for solved problems, for clear distinction of '.' and '..' for example -- amending lapse introduced in 9de4d64eee3b (April 2004);
wenzelm
parents: 39313
diff changeset
   825
code_pred divmod_rel .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   826
thm divmod_rel.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   827
value [code] "Predicate.the (divmod_rel_i_i_o_o 1705 42)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   828
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   829
subsection {* Transforming predicate logic into logic programs *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   830
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   831
subsection {* Transforming functions into logic programs *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   832
definition
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   833
  "case_f xs ys = (case (xs @ ys) of [] => [] | (x # xs) => xs)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   834
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
   835
code_pred [inductify, skip_proof] case_f .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   836
thm case_fP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   837
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   838
fun fold_map_idx where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   839
  "fold_map_idx f i y [] = (y, [])"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   840
| "fold_map_idx f i y (x # xs) =
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   841
 (let (y', x') = f i y x; (y'', xs') = fold_map_idx f (Suc i) y' xs
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   842
 in (y'', x' # xs'))"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   843
39195
5ab54bf226ac renewing specification in example file; adding invocation in example file
bulwahn
parents: 38664
diff changeset
   844
code_pred [inductify] fold_map_idx .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   845
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   846
subsection {* Minimum *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   847
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   848
definition Min
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   849
where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   850
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   851
code_pred [inductify] Min .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   852
thm Min.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   853
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   854
subsection {* Lexicographic order *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   855
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   856
declare lexord_def[code_pred_def]
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   857
code_pred [inductify] lexord .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   858
code_pred [random_dseq inductify] lexord .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   859
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   860
thm lexord.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   861
thm lexord.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   862
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   863
inductive less_than_nat :: "nat * nat => bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   864
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   865
  "less_than_nat (0, x)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   866
| "less_than_nat (x, y) ==> less_than_nat (Suc x, Suc y)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   867
 
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   868
code_pred less_than_nat .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   869
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   870
code_pred [dseq] less_than_nat .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   871
code_pred [random_dseq] less_than_nat .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   872
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   873
inductive test_lexord :: "nat list * nat list => bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   874
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   875
  "lexord less_than_nat (xs, ys) ==> test_lexord (xs, ys)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   876
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   877
code_pred test_lexord .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   878
code_pred [dseq] test_lexord .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   879
code_pred [random_dseq] test_lexord .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   880
thm test_lexord.dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   881
thm test_lexord.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   882
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   883
values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   884
(*values [depth_limited 5] "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   885
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   886
lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   887
(*
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   888
code_pred [inductify] lexn .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   889
thm lexn.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   890
*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   891
(*
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   892
code_pred [random_dseq inductify] lexn .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   893
thm lexn.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   894
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   895
values [random_dseq 4, 4, 6] 100 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   896
*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   897
inductive has_length
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   898
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   899
  "has_length [] 0"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   900
| "has_length xs i ==> has_length (x # xs) (Suc i)" 
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   901
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   902
lemma has_length:
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   903
  "has_length xs n = (length xs = n)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   904
proof (rule iffI)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   905
  assume "has_length xs n"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   906
  from this show "length xs = n"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   907
    by (rule has_length.induct) auto
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   908
next
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   909
  assume "length xs = n"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   910
  from this show "has_length xs n"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   911
    by (induct xs arbitrary: n) (auto intro: has_length.intros)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   912
qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   913
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   914
lemma lexn_intros [code_pred_intro]:
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   915
  "has_length xs i ==> has_length ys i ==> r (x, y) ==> lexn r (Suc i) (x # xs, y # ys)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   916
  "lexn r i (xs, ys) ==> lexn r (Suc i) (x # xs, x # ys)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   917
proof -
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   918
  assume "has_length xs i" "has_length ys i" "r (x, y)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   919
  from this has_length show "lexn r (Suc i) (x # xs, y # ys)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   920
    unfolding lexn_conv Collect_def mem_def
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   921
    by fastsimp
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   922
next
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   923
  assume "lexn r i (xs, ys)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   924
  thm lexn_conv
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   925
  from this show "lexn r (Suc i) (x#xs, x#ys)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   926
    unfolding Collect_def mem_def lexn_conv
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   927
    apply auto
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   928
    apply (rule_tac x="x # xys" in exI)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   929
    by auto
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   930
qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   931
36055
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
   932
code_pred [random_dseq] lexn
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   933
proof -
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   934
  fix r n xs ys
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   935
  assume 1: "lexn r n (xs, ys)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   936
  assume 2: "\<And>r' i x xs' y ys'. r = r' ==> n = Suc i ==> (xs, ys) = (x # xs', y # ys') ==> has_length xs' i ==> has_length ys' i ==> r' (x, y) ==> thesis"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   937
  assume 3: "\<And>r' i x xs' ys'. r = r' ==> n = Suc i ==> (xs, ys) = (x # xs', x # ys') ==> lexn r' i (xs', ys') ==> thesis"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   938
  from 1 2 3 show thesis
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   939
    unfolding lexn_conv Collect_def mem_def
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   940
    apply (auto simp add: has_length)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   941
    apply (case_tac xys)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   942
    apply auto
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   943
    apply fastsimp
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   944
    apply fastsimp done
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   945
qed
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   946
36055
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
   947
values [random_dseq 1, 2, 5] 10 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   948
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
   949
code_pred [inductify, skip_proof] lex .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   950
thm lex.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   951
thm lex_def
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   952
declare lenlex_conv[code_pred_def]
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
   953
code_pred [inductify, skip_proof] lenlex .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   954
thm lenlex.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   955
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   956
code_pred [random_dseq inductify] lenlex .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   957
thm lenlex.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   958
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   959
values [random_dseq 4, 2, 4] 100 "{(xs, ys::int list). lenlex (%(x, y). x <= y) (xs, ys)}"
36055
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
   960
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   961
thm lists.intros
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   962
code_pred [inductify] lists .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   963
thm lists.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   964
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   965
subsection {* AVL Tree *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   966
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   967
datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   968
fun height :: "'a tree => nat" where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   969
"height ET = 0"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   970
| "height (MKT x l r h) = max (height l) (height r) + 1"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   971
39195
5ab54bf226ac renewing specification in example file; adding invocation in example file
bulwahn
parents: 38664
diff changeset
   972
primrec avl :: "'a tree => bool"
5ab54bf226ac renewing specification in example file; adding invocation in example file
bulwahn
parents: 38664
diff changeset
   973
where
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   974
  "avl ET = True"
39195
5ab54bf226ac renewing specification in example file; adding invocation in example file
bulwahn
parents: 38664
diff changeset
   975
| "avl (MKT x l r h) = ((height l = height r \<or> height l = 1 + height r \<or> height r = 1+height l) \<and> 
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   976
  h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   977
(*
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   978
code_pred [inductify] avl .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   979
thm avl.equation*)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   980
36257
3f3e9f18f302 adopting examples to changes in the predicate compiler
bulwahn
parents: 36253
diff changeset
   981
code_pred [new_random_dseq inductify] avl .
3f3e9f18f302 adopting examples to changes in the predicate compiler
bulwahn
parents: 36253
diff changeset
   982
thm avl.new_random_dseq_equation
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   983
36257
3f3e9f18f302 adopting examples to changes in the predicate compiler
bulwahn
parents: 36253
diff changeset
   984
values [new_random_dseq 2, 1, 7] 5 "{t:: int tree. avl t}"
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   985
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   986
fun set_of
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   987
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   988
"set_of ET = {}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   989
| "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   990
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   991
fun is_ord :: "nat tree => bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   992
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   993
"is_ord ET = True"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   994
| "is_ord (MKT n l r h) =
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   995
 ((\<forall>n' \<in> set_of l. n' < n) \<and> (\<forall>n' \<in> set_of r. n < n') \<and> is_ord l \<and> is_ord r)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   996
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   997
code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] set_of .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   998
thm set_of.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
   999
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1000
code_pred (expected_modes: i => bool) [inductify] is_ord .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1001
thm is_ord_aux.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1002
thm is_ord.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1003
36055
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1004
subsection {* Definitions about Relations *}
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1005
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1006
term "converse"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1007
code_pred (modes:
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1008
  (i * i => bool) => i * i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1009
  (i * o => bool) => o * i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1010
  (i * o => bool) => i * i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1011
  (o * i => bool) => i * o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1012
  (o * i => bool) => i * i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1013
  (o * o => bool) => o * o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1014
  (o * o => bool) => i * o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1015
  (o * o => bool) => o * i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1016
  (o * o => bool) => i * i => bool) [inductify] converse .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1017
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1018
thm converse.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1019
code_pred [inductify] rel_comp .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1020
thm rel_comp.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1021
code_pred [inductify] Image .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1022
thm Image.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1023
declare singleton_iff[code_pred_inline]
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1024
declare Id_on_def[unfolded Bex_def UNION_def singleton_iff, code_pred_def]
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1025
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1026
code_pred (expected_modes:
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1027
  (o => bool) => o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1028
  (o => bool) => i * o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1029
  (o => bool) => o * i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1030
  (o => bool) => i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1031
  (i => bool) => i * o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1032
  (i => bool) => o * i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1033
  (i => bool) => i => bool) [inductify] Id_on .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1034
thm Id_on.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1035
thm Domain_def
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1036
code_pred (modes:
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1037
  (o * o => bool) => o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1038
  (o * o => bool) => i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1039
  (i * o => bool) => i => bool) [inductify] Domain .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1040
thm Domain.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1041
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1042
thm Range_def
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1043
code_pred (modes:
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1044
  (o * o => bool) => o => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1045
  (o * o => bool) => i => bool,
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1046
  (o * i => bool) => i => bool) [inductify] Range .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1047
thm Range.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1048
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1049
code_pred [inductify] Field .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1050
thm Field.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1051
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1052
thm refl_on_def
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1053
code_pred [inductify] refl_on .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1054
thm refl_on.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1055
code_pred [inductify] total_on .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1056
thm total_on.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1057
code_pred [inductify] antisym .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1058
thm antisym.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1059
code_pred [inductify] trans .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1060
thm trans.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1061
code_pred [inductify] single_valued .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1062
thm single_valued.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1063
thm inv_image_def
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1064
code_pred [inductify] inv_image .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1065
thm inv_image.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1066
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1067
subsection {* Inverting list functions *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1068
36055
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1069
code_pred [inductify] size_list .
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1070
code_pred [new_random_dseq inductify] size_list .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1071
thm size_listP.equation
36055
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1072
thm size_listP.new_random_dseq_equation
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1073
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1074
values [new_random_dseq 2,3,10] 3 "{xs. size_listP (xs::nat list) (5::nat)}"
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1075
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1076
code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) [inductify, skip_proof] List.concat .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1077
thm concatP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1078
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1079
values "{ys. concatP [[1, 2], [3, (4::int)]] ys}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1080
values "{ys. concatP [[1, 2], [3]] [1, 2, (3::nat)]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1081
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1082
code_pred [dseq inductify] List.concat .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1083
thm concatP.dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1084
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1085
values [dseq 3] 3
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1086
  "{xs. concatP xs ([0] :: nat list)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1087
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1088
values [dseq 5] 3
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1089
  "{xs. concatP xs ([1] :: int list)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1090
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1091
values [dseq 5] 3
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1092
  "{xs. concatP xs ([1] :: nat list)}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1093
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1094
values [dseq 5] 3
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1095
  "{xs. concatP xs [(1::int), 2]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1096
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1097
code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] hd .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1098
thm hdP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1099
values "{x. hdP [1, 2, (3::int)] x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1100
values "{(xs, x). hdP [1, 2, (3::int)] 1}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1101
 
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1102
code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] tl .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1103
thm tlP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1104
values "{x. tlP [1, 2, (3::nat)] x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1105
values "{x. tlP [1, 2, (3::int)] [3]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1106
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1107
code_pred [inductify, skip_proof] last .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1108
thm lastP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1109
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1110
code_pred [inductify, skip_proof] butlast .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1111
thm butlastP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1112
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1113
code_pred [inductify, skip_proof] take .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1114
thm takeP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1115
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1116
code_pred [inductify, skip_proof] drop .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1117
thm dropP.equation
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1118
code_pred [inductify, skip_proof] zip .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1119
thm zipP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1120
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1121
code_pred [inductify, skip_proof] upt .
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1122
code_pred [inductify, skip_proof] remdups .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1123
thm remdupsP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1124
code_pred [dseq inductify] remdups .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1125
values [dseq 4] 5 "{xs. remdupsP xs [1, (2::int)]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1126
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1127
code_pred [inductify, skip_proof] remove1 .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1128
thm remove1P.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1129
values "{xs. remove1P 1 xs [2, (3::int)]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1130
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1131
code_pred [inductify, skip_proof] removeAll .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1132
thm removeAllP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1133
code_pred [dseq inductify] removeAll .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1134
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1135
values [dseq 4] 10 "{xs. removeAllP 1 xs [(2::nat)]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1136
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1137
code_pred [inductify] distinct .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1138
thm distinct.equation
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1139
code_pred [inductify, skip_proof] replicate .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1140
thm replicateP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1141
values 5 "{(n, xs). replicateP n (0::int) xs}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1142
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1143
code_pred [inductify, skip_proof] splice .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1144
thm splice.simps
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1145
thm spliceP.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1146
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1147
values "{xs. spliceP xs [1, 2, 3] [1, 1, 1, 2, 1, (3::nat)]}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1148
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1149
code_pred [inductify, skip_proof] List.rev .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1150
code_pred [inductify] map .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1151
code_pred [inductify] foldr .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1152
code_pred [inductify] foldl .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1153
code_pred [inductify] filter .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1154
code_pred [random_dseq inductify] filter .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1155
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1156
subsection {* Context Free Grammar *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1157
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1158
datatype alphabet = a | b
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1159
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1160
inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1161
  "[] \<in> S\<^isub>1"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1162
| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1163
| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1164
| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1165
| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1166
| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1167
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1168
code_pred [inductify] S\<^isub>1p .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1169
code_pred [random_dseq inductify] S\<^isub>1p .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1170
thm S\<^isub>1p.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1171
thm S\<^isub>1p.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1172
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1173
values [random_dseq 5, 5, 5] 5 "{x. S\<^isub>1p x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1174
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1175
inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1176
  "[] \<in> S\<^isub>2"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1177
| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1178
| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1179
| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1180
| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1181
| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1182
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1183
code_pred [random_dseq inductify] S\<^isub>2p .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1184
thm S\<^isub>2p.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1185
thm A\<^isub>2p.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1186
thm B\<^isub>2p.random_dseq_equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1187
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1188
values [random_dseq 5, 5, 5] 10 "{x. S\<^isub>2p x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1189
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1190
inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1191
  "[] \<in> S\<^isub>3"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1192
| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1193
| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1194
| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1195
| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1196
| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1197
36040
fcd7bea01a93 adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
bulwahn
parents: 35954
diff changeset
  1198
code_pred [inductify, skip_proof] S\<^isub>3p .
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1199
thm S\<^isub>3p.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1200
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1201
values 10 "{x. S\<^isub>3p x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1202
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1203
inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1204
  "[] \<in> S\<^isub>4"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1205
| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1206
| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1207
| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1208
| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1209
| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1210
| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1211
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1212
code_pred (expected_modes: o => bool, i => bool) S\<^isub>4p .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1213
36176
3fe7e97ccca8 replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents: 36055
diff changeset
  1214
hide_const a b
36055
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1215
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1216
subsection {* Lambda *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1217
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1218
datatype type =
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1219
    Atom nat
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1220
  | Fun type type    (infixr "\<Rightarrow>" 200)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1221
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1222
datatype dB =
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1223
    Var nat
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1224
  | App dB dB (infixl "\<degree>" 200)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1225
  | Abs type dB
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1226
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1227
primrec
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1228
  nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1229
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1230
  "[]\<langle>i\<rangle> = None"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1231
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1232
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1233
inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1234
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1235
  "nth_el' (x # xs) 0 x"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1236
| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1237
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1238
inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile> _ : _" [50, 50, 50] 50)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1239
  where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1240
    Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1241
  | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1242
  | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1243
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1244
primrec
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1245
  lift :: "[dB, nat] => dB"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1246
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1247
    "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1248
  | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1249
  | "lift (Abs T s) k = Abs T (lift s (k + 1))"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1250
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1251
primrec
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1252
  subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1253
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1254
    subst_Var: "(Var i)[s/k] =
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1255
      (if k < i then Var (i - 1) else if i = k then s else Var i)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1256
  | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1257
  | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1258
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1259
inductive beta :: "[dB, dB] => bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1260
  where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1261
    beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1262
  | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1263
  | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1264
  | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1265
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1266
code_pred (expected_modes: i => i => o => bool, i => i => i => bool) typing .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1267
thm typing.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1268
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1269
code_pred (modes: i => i => bool,  i => o => bool as reduce') beta .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1270
thm beta.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1271
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1272
values "{x. App (Abs (Atom 0) (Var 0)) (Var 1) \<rightarrow>\<^sub>\<beta> x}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1273
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1274
definition "reduce t = Predicate.the (reduce' t)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1275
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1276
value "reduce (App (Abs (Atom 0) (Var 0)) (Var 1))"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1277
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1278
code_pred [dseq] typing .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1279
code_pred [random_dseq] typing .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1280
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1281
values [random_dseq 1,1,5] 10 "{(\<Gamma>, t, T). \<Gamma> \<turnstile> t : T}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1282
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1283
subsection {* A minimal example of yet another semantics *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1284
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1285
text {* thanks to Elke Salecker *}
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1286
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1287
types
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1288
  vname = nat
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1289
  vvalue = int
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1290
  var_assign = "vname \<Rightarrow> vvalue"  --"variable assignment"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1291
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1292
datatype ir_expr = 
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1293
  IrConst vvalue
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1294
| ObjAddr vname
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1295
| Add ir_expr ir_expr
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1296
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1297
datatype val =
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1298
  IntVal  vvalue
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1299
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1300
record  configuration =
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1301
  Env :: var_assign
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1302
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1303
inductive eval_var ::
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1304
  "ir_expr \<Rightarrow> configuration \<Rightarrow> val \<Rightarrow> bool"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1305
where
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1306
  irconst: "eval_var (IrConst i) conf (IntVal i)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1307
| objaddr: "\<lbrakk> Env conf n = i \<rbrakk> \<Longrightarrow> eval_var (ObjAddr n) conf (IntVal i)"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1308
| plus: "\<lbrakk> eval_var l conf (IntVal vl); eval_var r conf (IntVal vr) \<rbrakk> \<Longrightarrow> eval_var (Add l r) conf (IntVal (vl+vr))"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1309
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1310
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1311
code_pred eval_var .
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1312
thm eval_var.equation
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1313
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1314
values "{val. eval_var (Add (IrConst 1) (IrConst 2)) (| Env = (\<lambda>x. 0)|) val}"
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1315
36055
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1316
section {* Function predicate replacement *}
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1317
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1318
text {*
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1319
If the mode analysis uses the functional mode, we
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1320
replace predicates that resulted from functions again by their functions.
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1321
*}
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1322
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1323
inductive test_append
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1324
where
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1325
  "List.append xs ys = zs ==> test_append xs ys zs"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1326
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1327
code_pred [inductify, skip_proof] test_append .
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1328
thm test_append.equation
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1329
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1330
text {* If append is not turned into a predicate, then the mode
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1331
  o => o => i => bool could not be inferred. *}
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1332
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1333
values 4 "{(xs::int list, ys). test_append xs ys [1, 2, 3, 4]}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1334
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1335
text {* If appendP is not reverted back to a function, then mode i => i => o => bool
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1336
  fails after deleting the predicate equation. *}
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1337
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1338
declare appendP.equation[code del]
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1339
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1340
values "{xs::int list. test_append [1,2] [3,4] xs}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1341
values "{xs::int list. test_append (replicate 1000 1) (replicate 1000 2) xs}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1342
values "{xs::int list. test_append (replicate 2000 1) (replicate 2000 2) xs}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1343
39467
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1344
text {* Redeclaring append.equation as code equation *}
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1345
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1346
declare appendP.equation[code]
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1347
36055
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1348
subsection {* Function with tuples *}
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1349
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1350
fun append'
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1351
where
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1352
  "append' ([], ys) = ys"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1353
| "append' (x # xs, ys) = x # append' (xs, ys)"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1354
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1355
inductive test_append'
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1356
where
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1357
  "append' (xs, ys) = zs ==> test_append' xs ys zs"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1358
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1359
code_pred [inductify, skip_proof] test_append' .
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1360
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1361
thm test_append'.equation
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1362
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1363
values "{(xs::int list, ys). test_append' xs ys [1, 2, 3, 4]}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1364
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1365
declare append'P.equation[code del]
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1366
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1367
values "{zs :: int list. test_append' [1,2,3] [4,5] zs}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1368
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1369
section {* Arithmetic examples *}
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1370
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1371
subsection {* Examples on nat *}
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1372
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1373
inductive plus_nat_test :: "nat => nat => nat => bool"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1374
where
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1375
  "x + y = z ==> plus_nat_test x y z"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1376
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1377
code_pred [inductify, skip_proof] plus_nat_test .
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1378
code_pred [new_random_dseq inductify] plus_nat_test .
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1379
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1380
thm plus_nat_test.equation
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1381
thm plus_nat_test.new_random_dseq_equation
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1382
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1383
values [expected "{9::nat}"] "{z. plus_nat_test 4 5 z}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1384
values [expected "{9::nat}"] "{z. plus_nat_test 7 2 z}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1385
values [expected "{4::nat}"] "{y. plus_nat_test 5 y 9}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1386
values [expected "{}"] "{y. plus_nat_test 9 y 8}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1387
values [expected "{6::nat}"] "{y. plus_nat_test 1 y 7}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1388
values [expected "{2::nat}"] "{x. plus_nat_test x 7 9}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1389
values [expected "{}"] "{x. plus_nat_test x 9 7}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1390
values [expected "{(0::nat,0::nat)}"] "{(x, y). plus_nat_test x y 0}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1391
values [expected "{(0, Suc 0), (Suc 0, 0)}"] "{(x, y). plus_nat_test x y 1}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1392
values [expected "{(0, 5), (4, Suc 0), (3, 2), (2, 3), (Suc 0, 4), (5, 0)}"]
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1393
  "{(x, y). plus_nat_test x y 5}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1394
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1395
inductive minus_nat_test :: "nat => nat => nat => bool"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1396
where
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1397
  "x - y = z ==> minus_nat_test x y z"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1398
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1399
code_pred [inductify, skip_proof] minus_nat_test .
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1400
code_pred [new_random_dseq inductify] minus_nat_test .
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1401
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1402
thm minus_nat_test.equation
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1403
thm minus_nat_test.new_random_dseq_equation
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1404
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1405
values [expected "{0::nat}"] "{z. minus_nat_test 4 5 z}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1406
values [expected "{5::nat}"] "{z. minus_nat_test 7 2 z}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1407
values [expected "{16::nat}"] "{x. minus_nat_test x 7 9}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1408
values [expected "{16::nat}"] "{x. minus_nat_test x 9 7}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1409
values [expected "{0, Suc 0, 2, 3}"] "{x. minus_nat_test x 3 0}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1410
values [expected "{0::nat}"] "{x. minus_nat_test x 0 0}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1411
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1412
subsection {* Examples on int *}
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1413
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1414
inductive plus_int_test :: "int => int => int => bool"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1415
where
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1416
  "a + b = c ==> plus_int_test a b c"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1417
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1418
code_pred [inductify, skip_proof] plus_int_test .
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1419
code_pred [new_random_dseq inductify] plus_int_test .
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1420
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1421
thm plus_int_test.equation
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1422
thm plus_int_test.new_random_dseq_equation
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1423
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1424
values [expected "{1::int}"] "{a. plus_int_test a 6 7}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1425
values [expected "{1::int}"] "{b. plus_int_test 6 b 7}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1426
values [expected "{11::int}"] "{c. plus_int_test 5 6 c}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1427
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1428
inductive minus_int_test :: "int => int => int => bool"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1429
where
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1430
  "a - b = c ==> minus_int_test a b c"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1431
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1432
code_pred [inductify, skip_proof] minus_int_test .
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1433
code_pred [new_random_dseq inductify] minus_int_test .
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1434
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1435
thm minus_int_test.equation
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1436
thm minus_int_test.new_random_dseq_equation
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1437
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1438
values [expected "{4::int}"] "{c. minus_int_test 9 5 c}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1439
values [expected "{9::int}"] "{a. minus_int_test a 4 5}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1440
values [expected "{- 1::int}"] "{b. minus_int_test 4 b 5}"
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1441
36253
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1442
subsection {* minus on bool *}
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1443
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1444
inductive All :: "nat => bool"
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1445
where
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1446
  "All x"
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1447
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1448
inductive None :: "nat => bool"
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1449
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1450
definition "test_minus_bool x = (None x - All x)"
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1451
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1452
code_pred [inductify] test_minus_bool .
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1453
thm test_minus_bool.equation
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1454
6e969ce3dfcc added further inlining of boolean constants to the predicate compiler
bulwahn
parents: 36176
diff changeset
  1455
values "{x. test_minus_bool x}"
36055
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1456
39254
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1457
subsection {* Functions *}
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1458
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1459
fun partial_hd :: "'a list => 'a option"
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1460
where
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1461
  "partial_hd [] = Option.None"
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1462
| "partial_hd (x # xs) = Some x"
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1463
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1464
inductive hd_predicate
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1465
where
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1466
  "partial_hd xs = Some x ==> hd_predicate xs x"
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1467
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1468
code_pred (expected_modes: i => i => bool, i => o => bool) hd_predicate .
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1469
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1470
thm hd_predicate.equation
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1471
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1472
subsection {* Locales *}
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1473
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1474
inductive hd_predicate2 :: "('a list => 'a option) => 'a list => 'a => bool"
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1475
where
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1476
  "partial_hd' xs = Some x ==> hd_predicate2 partial_hd' xs x"
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1477
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1478
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1479
thm hd_predicate2.intros
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1480
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1481
code_pred (expected_modes: i => i => i => bool, i => i => o => bool) hd_predicate2 .
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1482
thm hd_predicate2.equation
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1483
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1484
locale A = fixes partial_hd :: "'a list => 'a option" begin
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1485
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1486
inductive hd_predicate_in_locale :: "'a list => 'a => bool"
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1487
where
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1488
  "partial_hd xs = Some x ==> hd_predicate_in_locale xs x"
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1489
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1490
end
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1491
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1492
text {* The global introduction rules must be redeclared as introduction rules and then 
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1493
  one could invoke code_pred. *}
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1494
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1495
declare A.hd_predicate_in_locale.intros [unfolded Predicate.eq_is_eq[symmetric], code_pred_intro]
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1496
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1497
code_pred (expected_modes: i => i => i => bool, i => i => o => bool) A.hd_predicate_in_locale
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1498
unfolding eq_is_eq by (auto elim: A.hd_predicate_in_locale.cases)
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1499
    
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1500
interpretation A partial_hd .
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1501
thm hd_predicate_in_locale.intros
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1502
text {* A locally compliant solution with a trivial interpretation fails, because
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1503
the predicate compiler has very strict assumptions about the terms and their structure. *}
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1504
 
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1505
(*code_pred hd_predicate_in_locale .*)
344d97e7b342 adding an example to show how code_pred must be invoked with locales
bulwahn
parents: 39195
diff changeset
  1506
39255
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1507
section {* Integer example *}
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1508
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1509
definition three :: int
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1510
where "three = 3"
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1511
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1512
inductive is_three
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1513
where
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1514
  "is_three three"
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1515
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1516
code_pred is_three .
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1517
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1518
thm is_three.equation
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1519
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1520
section {* String.literal example *}
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1521
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1522
definition Error_1
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1523
where
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1524
  "Error_1 = STR ''Error 1''"
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1525
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1526
definition Error_2
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1527
where
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1528
  "Error_2 = STR ''Error 2''"
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1529
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1530
inductive "is_error" :: "String.literal \<Rightarrow> bool"
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1531
where
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1532
  "is_error Error_1"
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1533
| "is_error Error_2"
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1534
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1535
code_pred is_error .
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1536
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1537
thm is_error.equation
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1538
39275
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1539
inductive is_error' :: "String.literal \<Rightarrow> bool"
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1540
where
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1541
  "is_error' (STR ''Error1'')"
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1542
| "is_error' (STR ''Error2'')"
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1543
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1544
code_pred is_error' .
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1545
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1546
thm is_error'.equation
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1547
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1548
datatype ErrorObject = Error String.literal int
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1549
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1550
inductive is_error'' :: "ErrorObject \<Rightarrow> bool"
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1551
where
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1552
  "is_error'' (Error Error_1 3)"
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1553
| "is_error'' (Error Error_2 4)"
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1554
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1555
code_pred is_error'' .
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1556
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1557
thm is_error''.equation
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1558
39313
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1559
section {* Another function example *}
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1560
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1561
consts f :: "'a \<Rightarrow> 'a"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1562
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1563
inductive fun_upd :: "(('a * 'b) * ('a \<Rightarrow> 'b)) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1564
where
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1565
  "fun_upd ((x, a), s) (s(x := f a))"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1566
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1567
code_pred fun_upd .
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1568
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1569
thm fun_upd.equation
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1570
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1571
section {* Another semantics *}
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1572
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1573
types
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1574
  name = nat --"For simplicity in examples"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1575
  state' = "name \<Rightarrow> nat"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1576
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1577
datatype aexp = N nat | V name | Plus aexp aexp
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1578
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1579
fun aval :: "aexp \<Rightarrow> state' \<Rightarrow> nat" where
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1580
"aval (N n) _ = n" |
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1581
"aval (V x) st = st x" |
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1582
"aval (Plus e\<^isub>1 e\<^isub>2) st = aval e\<^isub>1 st + aval e\<^isub>2 st"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1583
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1584
datatype bexp = B bool | Not bexp | And bexp bexp | Less aexp aexp
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1585
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1586
primrec bval :: "bexp \<Rightarrow> state' \<Rightarrow> bool" where
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1587
"bval (B b) _ = b" |
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1588
"bval (Not b) st = (\<not> bval b st)" |
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1589
"bval (And b1 b2) st = (bval b1 st \<and> bval b2 st)" |
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1590
"bval (Less a\<^isub>1 a\<^isub>2) st = (aval a\<^isub>1 st < aval a\<^isub>2 st)"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1591
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1592
datatype
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1593
  com' = SKIP 
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1594
      | Assign name aexp         ("_ ::= _" [1000, 61] 61)
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1595
      | Semi   com'  com'          ("_; _"  [60, 61] 60)
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1596
      | If     bexp com' com'     ("IF _ THEN _ ELSE _"  [0, 0, 61] 61)
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1597
      | While  bexp com'         ("WHILE _ DO _"  [0, 61] 61)
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1598
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1599
inductive
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1600
  big_step :: "com' * state' \<Rightarrow> state' \<Rightarrow> bool" (infix "\<Rightarrow>" 55)
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1601
where
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1602
  Skip:    "(SKIP,s) \<Rightarrow> s"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1603
| Assign:  "(x ::= a,s) \<Rightarrow> s(x := aval a s)"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1604
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1605
| Semi:    "(c\<^isub>1,s\<^isub>1) \<Rightarrow> s\<^isub>2  \<Longrightarrow>  (c\<^isub>2,s\<^isub>2) \<Rightarrow> s\<^isub>3  \<Longrightarrow> (c\<^isub>1;c\<^isub>2, s\<^isub>1) \<Rightarrow> s\<^isub>3"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1606
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1607
| IfTrue:  "bval b s  \<Longrightarrow>  (c\<^isub>1,s) \<Rightarrow> t  \<Longrightarrow>  (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1608
| IfFalse: "\<not>bval b s  \<Longrightarrow>  (c\<^isub>2,s) \<Rightarrow> t  \<Longrightarrow>  (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1609
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1610
| WhileFalse: "\<not>bval b s \<Longrightarrow> (WHILE b DO c,s) \<Rightarrow> s"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1611
| WhileTrue:  "bval b s\<^isub>1  \<Longrightarrow>  (c,s\<^isub>1) \<Rightarrow> s\<^isub>2  \<Longrightarrow>  (WHILE b DO c, s\<^isub>2) \<Rightarrow> s\<^isub>3
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1612
               \<Longrightarrow> (WHILE b DO c, s\<^isub>1) \<Rightarrow> s\<^isub>3"
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1613
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1614
code_pred big_step .
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1615
41ce0b56d858 adding two more examples to example theory
bulwahn
parents: 39275
diff changeset
  1616
thm big_step.equation
39275
dd84daec5d3c adding another String.literal example
bulwahn
parents: 39255
diff changeset
  1617
39467
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1618
section {* General Context Free Grammars *}
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1619
text {* a contribution by Aditi Barthwal *}
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1620
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1621
datatype ('nts,'ts) symbol = NTS 'nts
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1622
                            | TS 'ts
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1623
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1624
                            
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1625
datatype ('nts,'ts) rule = rule 'nts "('nts,'ts) symbol list"
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1626
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1627
types ('nts,'ts) grammar = "('nts,'ts) rule set * 'nts"
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1628
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1629
fun rules :: "('nts,'ts) grammar => ('nts,'ts) rule set"
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1630
where
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1631
  "rules (r, s) = r"
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1632
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1633
definition derives 
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1634
where
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1635
"derives g = { (lsl,rsl). \<exists>s1 s2 lhs rhs. 
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1636
                         (s1 @ [NTS lhs] @ s2 = lsl) \<and>
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1637
                         (s1 @ rhs @ s2) = rsl \<and>
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1638
                         (rule lhs rhs) \<in> fst g }"
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1639
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1640
abbreviation "example_grammar == 
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1641
({ rule ''S'' [NTS ''A'', NTS ''B''],
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1642
   rule ''S'' [TS ''a''],
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1643
  rule ''A'' [TS ''b'']}, ''S'')"
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1644
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1645
39543
9ff9651757cd removing unnessary options for code_pred
bulwahn
parents: 39506
diff changeset
  1646
code_pred [inductify, skip_proof] derives .
39467
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1647
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1648
thm derives.equation
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1649
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1650
definition "test = { rhs. derives example_grammar ([NTS ''S''], rhs) }"
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1651
39543
9ff9651757cd removing unnessary options for code_pred
bulwahn
parents: 39506
diff changeset
  1652
code_pred (modes: o \<Rightarrow> bool) [inductify] test .
39467
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1653
thm test.equation
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1654
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1655
values "{rhs. test rhs}"
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1656
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1657
declare rtrancl.intros(1)[code_pred_def] converse_rtrancl_into_rtrancl[code_pred_def]
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1658
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1659
code_pred [inductify] rtrancl .
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1660
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1661
definition "test2 = { rhs. ([NTS ''S''],rhs) \<in> (derives example_grammar)^*  }"
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1662
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1663
code_pred [inductify, skip_proof] test2 .
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1664
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1665
values "{rhs. test2 rhs}"
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1666
139c694299f6 adding another context free grammar example for the predicate compiler
bulwahn
parents: 39313
diff changeset
  1667
39255
432ed333294c adding an example with integers and String.literals
bulwahn
parents: 39254
diff changeset
  1668
section {* Examples for detecting switches *}
36260
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1669
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1670
inductive detect_switches1 where
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1671
  "detect_switches1 [] []"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1672
| "detect_switches1 (x # xs) (y # ys)"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1673
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1674
code_pred [detect_switches, skip_proof] detect_switches1 .
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1675
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1676
thm detect_switches1.equation
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1677
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1678
inductive detect_switches2 :: "('a => bool) => bool"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1679
where
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1680
  "detect_switches2 P"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1681
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1682
code_pred [detect_switches, skip_proof] detect_switches2 .
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1683
thm detect_switches2.equation
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1684
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1685
inductive detect_switches3 :: "('a => bool) => 'a list => bool"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1686
where
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1687
  "detect_switches3 P []"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1688
| "detect_switches3 P (x # xs)" 
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1689
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1690
code_pred [detect_switches, skip_proof] detect_switches3 .
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1691
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1692
thm detect_switches3.equation
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1693
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1694
inductive detect_switches4 :: "('a => bool) => 'a list => 'a list => bool"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1695
where
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1696
  "detect_switches4 P [] []"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1697
| "detect_switches4 P (x # xs) (y # ys)"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1698
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1699
code_pred [detect_switches, skip_proof] detect_switches4 .
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1700
thm detect_switches4.equation
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1701
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1702
inductive detect_switches5 :: "('a => 'a => bool) => 'a list => 'a list => bool"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1703
where
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1704
  "detect_switches5 P [] []"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1705
| "detect_switches5 P xs ys ==> P x y ==> detect_switches5 P (x # xs) (y # ys)"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1706
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1707
code_pred [detect_switches, skip_proof] detect_switches5 .
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1708
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1709
thm detect_switches5.equation
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1710
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1711
inductive detect_switches6 :: "(('a => 'b => bool) * 'a list * 'b list) => bool"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1712
where
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1713
  "detect_switches6 (P, [], [])"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1714
| "detect_switches6 (P, xs, ys) ==> P x y ==> detect_switches6 (P, x # xs, y # ys)"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1715
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1716
code_pred [detect_switches, skip_proof] detect_switches6 .
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1717
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1718
inductive detect_switches7 :: "('a => bool) => ('b => bool) => ('a * 'b list) => bool"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1719
where
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1720
  "detect_switches7 P Q (a, [])"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1721
| "P a ==> Q x ==> detect_switches7 P Q (a, x#xs)"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1722
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1723
code_pred [skip_proof] detect_switches7 .
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1724
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1725
thm detect_switches7.equation
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1726
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1727
inductive detect_switches8 :: "nat => bool"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1728
where
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1729
  "detect_switches8 0"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1730
| "x mod 2 = 0 ==> detect_switches8 (Suc x)"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1731
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1732
code_pred [detect_switches, skip_proof] detect_switches8 .
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1733
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1734
thm detect_switches8.equation
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1735
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1736
inductive detect_switches9 :: "nat => nat => bool"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1737
where
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1738
  "detect_switches9 0 0"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1739
| "detect_switches9 0 (Suc x)"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1740
| "detect_switches9 (Suc x) 0"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1741
| "x = y ==> detect_switches9 (Suc x) (Suc y)"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1742
| "c1 = c2 ==> detect_switches9 c1 c2"
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1743
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1744
code_pred [detect_switches, skip_proof] detect_switches9 .
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1745
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1746
thm detect_switches9.equation
c0ab79a100e4 added examples for detecting switches
bulwahn
parents: 36257
diff changeset
  1747
36055
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1748
537876d0fa62 adding examples of function predicate replacement and arithmetic examples for the predicate compiler; tuned
bulwahn
parents: 36040
diff changeset
  1749
35950
791ce568d40a moved examples for the predicate compiler into its own session Predicate_Compile_Examples to slenderise the HOL-ex session
bulwahn
parents:
diff changeset
  1750
end