src/HOL/ex/Predicate_Compile_ex.thy
author haftmann
Wed, 10 Feb 2010 08:49:26 +0100
changeset 35084 e25eedfc15ce
parent 34948 2d5f2a9f7601
child 35324 c9f428269b38
permissions -rw-r--r--
moved constants inverse and divide to Ring.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31129
d2cead76fca2 split Predicate_Compile examples into separate theory
haftmann
parents: 31123
diff changeset
     1
theory Predicate_Compile_ex
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
     2
imports Predicate_Compile_Alternative_Defs
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     3
begin
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     4
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
     5
subsection {* Basic predicates *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
     6
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
     7
inductive False' :: "bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
     8
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
     9
code_pred (expected_modes: bool) False' .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    10
code_pred [dseq] False' .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    11
code_pred [random_dseq] False' .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    12
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    13
values [expected "{}" pred] "{x. False'}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    14
values [expected "{}" dseq 1] "{x. False'}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    15
values [expected "{}" random_dseq 1, 1, 1] "{x. False'}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    16
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    17
value "False'"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    18
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    19
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    20
inductive True' :: "bool"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    21
where
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    22
  "True ==> True'"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    23
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    24
code_pred True' .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    25
code_pred [dseq] True' .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    26
code_pred [random_dseq] True' .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    27
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    28
thm True'.equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    29
thm True'.dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    30
thm True'.random_dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    31
values [expected "{()}" ]"{x. True'}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    32
values [expected "{}" dseq 0] "{x. True'}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    33
values [expected "{()}" dseq 1] "{x. True'}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    34
values [expected "{()}" dseq 2] "{x. True'}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    35
values [expected "{}" random_dseq 1, 1, 0] "{x. True'}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    36
values [expected "{}" random_dseq 1, 1, 1] "{x. True'}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    37
values [expected "{()}" random_dseq 1, 1, 2] "{x. True'}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    38
values [expected "{()}" random_dseq 1, 1, 3] "{x. True'}"
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    39
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    40
inductive EmptySet :: "'a \<Rightarrow> bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    41
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
    42
code_pred (expected_modes: o => bool, i => bool) EmptySet .
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    43
33253
d9ca0c3bf680 changes to example file
bulwahn
parents: 33250
diff changeset
    44
definition EmptySet' :: "'a \<Rightarrow> bool"
d9ca0c3bf680 changes to example file
bulwahn
parents: 33250
diff changeset
    45
where "EmptySet' = {}"
d9ca0c3bf680 changes to example file
bulwahn
parents: 33250
diff changeset
    46
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
    47
code_pred (expected_modes: o => bool, i => bool) [inductify] EmptySet' .
33253
d9ca0c3bf680 changes to example file
bulwahn
parents: 33250
diff changeset
    48
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    49
inductive EmptyRel :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    50
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
    51
code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) EmptyRel .
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    52
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    53
inductive EmptyClosure :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    54
for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    55
33327
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
    56
code_pred
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
    57
  (expected_modes: (o => o => bool) => o => o => bool, (o => o => bool) => i => o => bool,
33621
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
    58
         (o => o => bool) => o => i => bool, (o => o => bool) => i => i => bool,
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
    59
         (i => o => bool) => o => o => bool, (i => o => bool) => i => o => bool,
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
    60
         (i => o => bool) => o => i => bool, (i => o => bool) => i => i => bool,
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
    61
         (o => i => bool) => o => o => bool, (o => i => bool) => i => o => bool,
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
    62
         (o => i => bool) => o => i => bool, (o => i => bool) => i => i => bool,
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
    63
         (i => i => bool) => o => o => bool, (i => i => bool) => i => o => bool,
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
    64
         (i => i => bool) => o => i => bool, (i => i => bool) => i => i => bool)
33327
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
    65
  EmptyClosure .
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
    66
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    67
thm EmptyClosure.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    68
33258
f47ca46d2187 disabled test case for predicate compiler due to an problem in the inductive package
bulwahn
parents: 33255
diff changeset
    69
(* TODO: inductive package is broken!
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    70
inductive False'' :: "bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    71
where
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    72
  "False \<Longrightarrow> False''"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    73
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
    74
code_pred (expected_modes: []) False'' .
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    75
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
    76
inductive EmptySet'' :: "'a \<Rightarrow> bool"
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    77
where
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
    78
  "False \<Longrightarrow> EmptySet'' x"
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    79
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
    80
code_pred (expected_modes: [1]) EmptySet'' .
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
    81
code_pred (expected_modes: [], [1]) [inductify] EmptySet'' .
33258
f47ca46d2187 disabled test case for predicate compiler due to an problem in the inductive package
bulwahn
parents: 33255
diff changeset
    82
*)
33621
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
    83
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    84
consts a' :: 'a
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    85
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    86
inductive Fact :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    87
where
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    88
"Fact a' a'"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    89
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
    90
code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) Fact .
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    91
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    92
inductive zerozero :: "nat * nat => bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    93
where
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    94
  "zerozero (0, 0)"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    95
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
    96
code_pred (expected_modes: i => bool, i * o => bool, o * i => bool, o => bool) zerozero .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    97
code_pred [dseq] zerozero .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    98
code_pred [random_dseq] zerozero .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
    99
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   100
thm zerozero.equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   101
thm zerozero.dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   102
thm zerozero.random_dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   103
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   104
text {* We expect the user to expand the tuples in the values command.
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   105
The following values command is not supported. *}
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   106
(*values "{x. zerozero x}" *)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   107
text {* Instead, the user must type *}
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   108
values "{(x, y). zerozero (x, y)}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   109
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   110
values [expected "{}" dseq 0] "{(x, y). zerozero (x, y)}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   111
values [expected "{(0::nat, 0::nat)}" dseq 1] "{(x, y). zerozero (x, y)}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   112
values [expected "{(0::nat, 0::nat)}" dseq 2] "{(x, y). zerozero (x, y)}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   113
values [expected "{}" random_dseq 1, 1, 2] "{(x, y). zerozero (x, y)}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   114
values [expected "{(0::nat, 0:: nat)}" random_dseq 1, 1, 3] "{(x, y). zerozero (x, y)}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   115
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   116
inductive nested_tuples :: "((int * int) * int * int) => bool"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   117
where
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   118
  "nested_tuples ((0, 1), 2, 3)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   119
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   120
code_pred nested_tuples .
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   121
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
   122
inductive JamesBond :: "nat => int => code_numeral => bool"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
   123
where
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
   124
  "JamesBond 0 0 7"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
   125
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
   126
code_pred JamesBond .
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
   127
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   128
values [expected "{(0::nat, 0::int , 7::code_numeral)}"] "{(a, b, c). JamesBond a b c}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   129
values [expected "{(0::nat, 7::code_numeral, 0:: int)}"] "{(a, c, b). JamesBond a b c}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   130
values [expected "{(0::int, 0::nat, 7::code_numeral)}"] "{(b, a, c). JamesBond a b c}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   131
values [expected "{(0::int, 7::code_numeral, 0::nat)}"] "{(b, c, a). JamesBond a b c}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   132
values [expected "{(7::code_numeral, 0::nat, 0::int)}"] "{(c, a, b). JamesBond a b c}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   133
values [expected "{(7::code_numeral, 0::int, 0::nat)}"] "{(c, b, a). JamesBond a b c}"
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
   134
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   135
values [expected "{(7::code_numeral, 0::int)}"] "{(a, b). JamesBond 0 b a}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   136
values [expected "{(7::code_numeral, 0::nat)}"] "{(c, a). JamesBond a 0 c}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   137
values [expected "{(0::nat, 7::code_numeral)}"] "{(a, c). JamesBond a 0 c}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   138
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
   139
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   140
subsection {* Alternative Rules *}
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   141
33329
b129e4c476d6 improved mode parser; added mode annotations to examples
bulwahn
parents: 33327
diff changeset
   142
datatype char = C | D | E | F | G | H
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   143
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   144
inductive is_C_or_D
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   145
where
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   146
  "(x = C) \<or> (x = D) ==> is_C_or_D x"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   147
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   148
code_pred (expected_modes: i => bool) is_C_or_D .
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   149
thm is_C_or_D.equation
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   150
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   151
inductive is_D_or_E
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   152
where
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   153
  "(x = D) \<or> (x = E) ==> is_D_or_E x"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   154
33628
ed2111a5c3ed renaming code_pred_intros to code_pred_intro
bulwahn
parents: 33626
diff changeset
   155
lemma [code_pred_intro]:
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   156
  "is_D_or_E D"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   157
by (auto intro: is_D_or_E.intros)
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   158
33628
ed2111a5c3ed renaming code_pred_intros to code_pred_intro
bulwahn
parents: 33626
diff changeset
   159
lemma [code_pred_intro]:
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   160
  "is_D_or_E E"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   161
by (auto intro: is_D_or_E.intros)
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   162
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   163
code_pred (expected_modes: o => bool, i => bool) is_D_or_E
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   164
proof -
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   165
  case is_D_or_E
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   166
  from this(1) show thesis
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   167
  proof
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   168
    fix xa
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   169
    assume x: "x = xa"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   170
    assume "xa = D \<or> xa = E"
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   171
    from this show thesis
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   172
    proof
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   173
      assume "xa = D" from this x is_D_or_E(2) show thesis by simp
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   174
    next
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   175
      assume "xa = E" from this x is_D_or_E(3) show thesis by simp
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   176
    qed
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   177
  qed
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   178
qed
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   179
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   180
thm is_D_or_E.equation
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   181
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   182
inductive is_F_or_G
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   183
where
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   184
  "x = F \<or> x = G ==> is_F_or_G x"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   185
33628
ed2111a5c3ed renaming code_pred_intros to code_pred_intro
bulwahn
parents: 33626
diff changeset
   186
lemma [code_pred_intro]:
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   187
  "is_F_or_G F"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   188
by (auto intro: is_F_or_G.intros)
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   189
33628
ed2111a5c3ed renaming code_pred_intros to code_pred_intro
bulwahn
parents: 33626
diff changeset
   190
lemma [code_pred_intro]:
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   191
  "is_F_or_G G"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   192
by (auto intro: is_F_or_G.intros)
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   193
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   194
inductive is_FGH
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   195
where
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   196
  "is_F_or_G x ==> is_FGH x"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   197
| "is_FGH H"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   198
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   199
text {* Compilation of is_FGH requires elimination rule for is_F_or_G *}
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   200
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   201
code_pred (expected_modes: o => bool, i => bool) is_FGH
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   202
proof -
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   203
  case is_F_or_G
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   204
  from this(1) show thesis
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   205
  proof
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   206
    fix xa
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   207
    assume x: "x = xa"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   208
    assume "xa = F \<or> xa = G"
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   209
    from this show thesis
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   210
    proof
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   211
      assume "xa = F"
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   212
      from this x is_F_or_G(2) show thesis by simp
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   213
    next
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   214
      assume "xa = G"
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   215
      from this x is_F_or_G(3) show thesis by simp
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   216
    qed
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   217
  qed
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   218
qed
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   219
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   220
subsection {* Preprocessor Inlining  *}
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   221
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   222
definition "equals == (op =)"
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   223
 
33405
5c1928d5db38 adding testcases for code inlining, sets and numerals to the example files of the predicate compiler and its quickcheck prototype; disabling length example
bulwahn
parents: 33375
diff changeset
   224
inductive zerozero' :: "nat * nat => bool" where
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   225
  "equals (x, y) (0, 0) ==> zerozero' (x, y)"
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   226
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   227
code_pred (expected_modes: i => bool) zerozero' .
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   228
33405
5c1928d5db38 adding testcases for code inlining, sets and numerals to the example files of the predicate compiler and its quickcheck prototype; disabling length example
bulwahn
parents: 33375
diff changeset
   229
lemma zerozero'_eq: "zerozero' x == zerozero x"
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   230
proof -
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   231
  have "zerozero' = zerozero"
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   232
    apply (auto simp add: mem_def)
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   233
    apply (cases rule: zerozero'.cases)
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   234
    apply (auto simp add: equals_def intro: zerozero.intros)
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   235
    apply (cases rule: zerozero.cases)
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   236
    apply (auto simp add: equals_def intro: zerozero'.intros)
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   237
    done
33405
5c1928d5db38 adding testcases for code inlining, sets and numerals to the example files of the predicate compiler and its quickcheck prototype; disabling length example
bulwahn
parents: 33375
diff changeset
   238
  from this show "zerozero' x == zerozero x" by auto
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   239
qed
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   240
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   241
declare zerozero'_eq [code_pred_inline]
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   242
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   243
definition "zerozero'' x == zerozero' x"
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   244
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   245
text {* if preprocessing fails, zerozero'' will not have all modes. *}
33405
5c1928d5db38 adding testcases for code inlining, sets and numerals to the example files of the predicate compiler and its quickcheck prototype; disabling length example
bulwahn
parents: 33375
diff changeset
   246
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   247
code_pred (expected_modes: i * i => bool, i * o => bool, o * i => bool, o => bool) [inductify] zerozero'' .
33405
5c1928d5db38 adding testcases for code inlining, sets and numerals to the example files of the predicate compiler and its quickcheck prototype; disabling length example
bulwahn
parents: 33375
diff changeset
   248
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   249
subsection {* Sets and Numerals *}
33405
5c1928d5db38 adding testcases for code inlining, sets and numerals to the example files of the predicate compiler and its quickcheck prototype; disabling length example
bulwahn
parents: 33375
diff changeset
   250
5c1928d5db38 adding testcases for code inlining, sets and numerals to the example files of the predicate compiler and its quickcheck prototype; disabling length example
bulwahn
parents: 33375
diff changeset
   251
definition
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   252
  "one_or_two = {Suc 0, (Suc (Suc 0))}"
33405
5c1928d5db38 adding testcases for code inlining, sets and numerals to the example files of the predicate compiler and its quickcheck prototype; disabling length example
bulwahn
parents: 33375
diff changeset
   253
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   254
code_pred [inductify] one_or_two .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   255
code_pred [dseq] one_or_two .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   256
(*code_pred [random_dseq] one_or_two .*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   257
values [expected "{Suc 0::nat, 2::nat}"] "{x. one_or_two x}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   258
(*values [random_dseq 1,1,2] "{x. one_or_two x}"*)
33405
5c1928d5db38 adding testcases for code inlining, sets and numerals to the example files of the predicate compiler and its quickcheck prototype; disabling length example
bulwahn
parents: 33375
diff changeset
   259
33475
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   260
inductive one_or_two' :: "nat => bool"
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   261
where
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   262
  "one_or_two' 1"
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   263
| "one_or_two' 2"
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   264
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   265
code_pred one_or_two' .
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   266
thm one_or_two'.equation
33405
5c1928d5db38 adding testcases for code inlining, sets and numerals to the example files of the predicate compiler and its quickcheck prototype; disabling length example
bulwahn
parents: 33375
diff changeset
   267
33475
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   268
values "{x. one_or_two' x}"
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   269
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   270
definition one_or_two'':
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   271
  "one_or_two'' == {1, (2::nat)}"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   272
ML {* prop_of @{thm one_or_two''} *}
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   273
(*code_pred [inductify] one_or_two'' .
33475
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   274
thm one_or_two''.equation
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   275
42fed8eac357 improved handling of overloaded constants; examples with numerals
bulwahn
parents: 33473
diff changeset
   276
values "{x. one_or_two'' x}"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   277
*)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   278
subsection {* even predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   279
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   280
inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   281
    "even 0"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   282
  | "even n \<Longrightarrow> odd (Suc n)"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   283
  | "odd n \<Longrightarrow> even (Suc n)"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   284
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   285
code_pred (expected_modes: i => bool, o => bool) even .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   286
code_pred [dseq] even .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   287
code_pred [random_dseq] even .
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33136
diff changeset
   288
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
   289
thm odd.equation
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   290
thm even.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   291
thm odd.dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   292
thm even.dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   293
thm odd.random_dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   294
thm even.random_dseq_equation
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   295
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   296
values "{x. even 2}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   297
values "{x. odd 2}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   298
values 10 "{n. even n}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   299
values 10 "{n. odd n}"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   300
values [expected "{}" dseq 2] "{x. even 6}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   301
values [expected "{}" dseq 6] "{x. even 6}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   302
values [expected "{()}" dseq 7] "{x. even 6}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   303
values [dseq 2] "{x. odd 7}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   304
values [dseq 6] "{x. odd 7}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   305
values [dseq 7] "{x. odd 7}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   306
values [expected "{()}" dseq 8] "{x. odd 7}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   307
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   308
values [expected "{}" dseq 0] 8 "{x. even x}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   309
values [expected "{0::nat}" dseq 1] 8 "{x. even x}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   310
values [expected "{0::nat, 2}" dseq 3] 8 "{x. even x}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   311
values [expected "{0::nat, 2}" dseq 4] 8 "{x. even x}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   312
values [expected "{0::nat, 2, 4}" dseq 6] 8 "{x. even x}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   313
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   314
values [random_dseq 1, 1, 0] 8 "{x. even x}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   315
values [random_dseq 1, 1, 1] 8 "{x. even x}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   316
values [random_dseq 1, 1, 2] 8 "{x. even x}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   317
values [random_dseq 1, 1, 3] 8 "{x. even x}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   318
values [random_dseq 1, 1, 6] 8 "{x. even x}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   319
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   320
values [expected "{}" random_dseq 1, 1, 7] "{x. odd 7}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   321
values [random_dseq 1, 1, 8] "{x. odd 7}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   322
values [random_dseq 1, 1, 9] "{x. odd 7}"
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   323
33136
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   324
definition odd' where "odd' x == \<not> even x"
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   325
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   326
code_pred (expected_modes: i => bool) [inductify] odd' .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   327
code_pred [dseq inductify] odd' .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   328
code_pred [random_dseq inductify] odd' .
33136
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   329
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   330
values [expected "{}" dseq 2] "{x. odd' 7}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   331
values [expected "{()}" dseq 9] "{x. odd' 7}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   332
values [expected "{}" dseq 2] "{x. odd' 8}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   333
values [expected "{}" dseq 10] "{x. odd' 8}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   334
33136
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   335
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   336
inductive is_even :: "nat \<Rightarrow> bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   337
where
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   338
  "n mod 2 = 0 \<Longrightarrow> is_even n"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   339
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   340
code_pred (expected_modes: i => bool) is_even .
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   341
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   342
subsection {* append predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   343
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   344
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   345
    "append [] xs xs"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   346
  | "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   347
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   348
code_pred (modes: i => i => o => bool as "concat", o => o => i => bool as "slice", o => i => i => bool as prefix,
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   349
  i => o => i => bool as suffix, i => i => i => bool) append .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   350
code_pred [dseq] append .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   351
code_pred [random_dseq] append .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   352
33477
1272cfc7b910 added further example of the values command
bulwahn
parents: 33475
diff changeset
   353
thm append.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   354
thm append.dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   355
thm append.random_dseq_equation
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   356
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   357
values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   358
values "{zs. append [0, Suc 0, 2] [17, 8] zs}"
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   359
values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0, 5]}"
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
   360
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   361
values [expected "{}" dseq 0] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   362
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)]}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   363
values [dseq 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   364
values [dseq 6] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   365
values [random_dseq 1, 1, 4] 10 "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   366
values [random_dseq 1, 1, 1] 10 "{(xs, ys, zs::int list). append xs ys zs}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   367
values [random_dseq 1, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   368
values [random_dseq 3, 1, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   369
values [random_dseq 1, 3, 3] 10 "{(xs, ys, zs::int list). append xs ys zs}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   370
values [random_dseq 1, 1, 4] 10 "{(xs, ys, zs::int list). append xs ys zs}"
31195
12741f23527d added example on ML level
haftmann
parents: 31129
diff changeset
   371
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   372
value [code] "Predicate.the (concat [0::int, 1, 2] [3, 4, 5])"
33621
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
   373
value [code] "Predicate.the (slice ([]::int list))"
33111
db5af7b86a2f developing an executable the operator
bulwahn
parents: 33108
diff changeset
   374
33480
dcfe9100e0a6 disabled upt example because of a problem due to overloaded constants with the predicate compiler
bulwahn
parents: 33479
diff changeset
   375
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   376
text {* tricky case with alternative rules *}
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   377
33116
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   378
inductive append2
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   379
where
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   380
  "append2 [] xs xs"
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   381
| "append2 xs ys zs \<Longrightarrow> append2 (x # xs) ys (x # zs)"
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   382
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   383
lemma append2_Nil: "append2 [] (xs::'b list) xs"
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   384
  by (simp add: append2.intros(1))
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   385
33628
ed2111a5c3ed renaming code_pred_intros to code_pred_intro
bulwahn
parents: 33626
diff changeset
   386
lemmas [code_pred_intro] = append2_Nil append2.intros(2)
33116
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   387
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   388
code_pred (expected_modes: i => i => o => bool, o => o => i => bool, o => i => i => bool,
33621
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
   389
  i => o => i => bool, i => i => i => bool) append2
33116
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   390
proof -
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   391
  case append2
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   392
  from append2(1) show thesis
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   393
  proof
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   394
    fix xs
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   395
    assume "xa = []" "xb = xs" "xc = xs"
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   396
    from this append2(2) show thesis by simp
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   397
  next
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   398
    fix xs ys zs x
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   399
    assume "xa = x # xs" "xb = ys" "xc = x # zs" "append2 xs ys zs"
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   400
    from this append2(3) show thesis by fastsimp
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   401
  qed
33116
b379ee2cddb1 changed importing introduction rules to fix the same type variables in all introduction rules; improved mode analysis for partially applied relations; added test case; tuned
bulwahn
parents: 33114
diff changeset
   402
qed
33128
1f990689349f further cleaning up
bulwahn
parents: 33126
diff changeset
   403
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   404
inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   405
where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   406
  "tupled_append ([], xs, xs)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   407
| "tupled_append (xs, ys, zs) \<Longrightarrow> tupled_append (x # xs, ys, x # zs)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   408
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   409
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
33621
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
   410
  i * o * i => bool, i * i * i => bool) tupled_append .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   411
code_pred [random_dseq] tupled_append .
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   412
thm tupled_append.equation
33621
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
   413
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   414
values "{xs. tupled_append ([(1::nat), 2, 3], [4, 5], xs)}"
33621
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
   415
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   416
inductive tupled_append'
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   417
where
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   418
"tupled_append' ([], xs, xs)"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   419
| "[| ys = fst (xa, y); x # zs = snd (xa, y);
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   420
 tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)"
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   421
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   422
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
33621
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
   423
  i * o * i => bool, i * i * i => bool) tupled_append' .
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   424
thm tupled_append'.equation
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   425
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   426
inductive tupled_append'' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   427
where
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   428
  "tupled_append'' ([], xs, xs)"
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   429
| "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, zs) \<Longrightarrow> tupled_append'' (x # xs, yszs)"
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   430
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   431
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   432
  i * o * i => bool, i * i * i => bool) tupled_append'' .
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   433
thm tupled_append''.equation
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   434
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   435
inductive tupled_append''' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   436
where
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   437
  "tupled_append''' ([], xs, xs)"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   438
| "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \<Longrightarrow> tupled_append''' (x # xs, ys, x # zs)"
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   439
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   440
code_pred (expected_modes: i * i * o => bool, o * o * i => bool, o * i * i => bool,
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   441
  i * o * i => bool, i * i * i => bool) tupled_append''' .
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   442
thm tupled_append'''.equation
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   443
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   444
subsection {* map_ofP predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   445
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   446
inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   447
where
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   448
  "map_ofP ((a, b)#xs) a b"
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   449
| "map_ofP xs a b \<Longrightarrow> map_ofP (x#xs) a b"
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   450
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   451
code_pred (expected_modes: i => o => o => bool, i => i => o => bool, i => o => i => bool, i => i => i => bool) map_ofP .
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   452
thm map_ofP.equation
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   453
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   454
subsection {* filter predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   455
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   456
inductive filter1
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   457
for P
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   458
where
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   459
  "filter1 P [] []"
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   460
| "P x ==> filter1 P xs ys ==> filter1 P (x#xs) (x#ys)"
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   461
| "\<not> P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys"
33329
b129e4c476d6 improved mode parser; added mode annotations to examples
bulwahn
parents: 33327
diff changeset
   462
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   463
code_pred (expected_modes: (i => bool) => i => o => bool, (i => bool) => i => i => bool) filter1 .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   464
code_pred [dseq] filter1 .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   465
code_pred [random_dseq] filter1 .
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   466
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   467
thm filter1.equation
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   468
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   469
values [expected "{[0::nat, 2, 4]}"] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   470
values [expected "{}" dseq 9] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   471
values [expected "{[0::nat, 2, 4]}" dseq 10] "{xs. filter1 even [0, 1, 2, 3, 4] xs}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   472
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   473
inductive filter2
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   474
where
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   475
  "filter2 P [] []"
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   476
| "P x ==> filter2 P xs ys ==> filter2 P (x#xs) (x#ys)"
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   477
| "\<not> P x ==> filter2 P xs ys ==> filter2 P (x#xs) ys"
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   478
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   479
code_pred (expected_modes: (i => bool) => i => i => bool, (i => bool) => i => o => bool) filter2 .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   480
code_pred [dseq] filter2 .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   481
code_pred [random_dseq] filter2 .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   482
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   483
thm filter2.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   484
thm filter2.random_dseq_equation
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   485
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   486
(*
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   487
inductive filter3
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   488
for P
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   489
where
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   490
  "List.filter P xs = ys ==> filter3 P xs ys"
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   491
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   492
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 .
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   493
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   494
code_pred [dseq] filter3 .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   495
thm filter3.dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   496
*)
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   497
inductive filter4
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   498
where
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   499
  "List.filter P xs = ys ==> filter4 P xs ys"
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   500
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   501
(*code_pred (expected_modes: i => i => o => bool, i => i => i => bool) filter4 .*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   502
(*code_pred [depth_limited] filter4 .*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   503
(*code_pred [random] filter4 .*)
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   504
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   505
subsection {* reverse predicate *}
33112
6672184a736b added a few tricky examples with tuples; added alternative introduction rules for some constants; corrected mode analysis with negation; improved fetching of definitions
bulwahn
parents: 33111
diff changeset
   506
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   507
inductive rev where
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   508
    "rev [] []"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   509
  | "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   510
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   511
code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) rev .
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   512
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   513
thm rev.equation
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   514
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   515
values "{xs. rev [0, 1, 2, 3::nat] xs}"
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   516
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   517
inductive tupled_rev where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   518
  "tupled_rev ([], [])"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   519
| "tupled_rev (xs, xs') \<Longrightarrow> tupled_append (xs', [x], ys) \<Longrightarrow> tupled_rev (x#xs, ys)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   520
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   521
code_pred (expected_modes: i * o => bool, o * i => bool, i * i => bool) tupled_rev .
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   522
thm tupled_rev.equation
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   523
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   524
subsection {* partition predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   525
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   526
inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   527
  for f where
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   528
    "partition f [] [] []"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   529
  | "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   530
  | "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   531
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   532
code_pred (expected_modes: (i => bool) => i => o => o => bool, (i => bool) => o => i => i => bool,
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   533
  (i => bool) => i => i => o => bool, (i => bool) => i => o => i => bool, (i => bool) => i => i => i => bool)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   534
  partition .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   535
code_pred [dseq] partition .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   536
code_pred [random_dseq] partition .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   537
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   538
values 10 "{(ys, zs). partition is_even
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   539
  [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   540
values 10 "{zs. partition is_even zs [0, 2] [3, 5]}"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   541
values 10 "{zs. partition is_even zs [0, 7] [3, 5]}"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   542
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   543
inductive tupled_partition :: "('a \<Rightarrow> bool) \<Rightarrow> ('a list \<times> 'a list \<times> 'a list) \<Rightarrow> bool"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   544
  for f where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   545
   "tupled_partition f ([], [], [])"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   546
  | "f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, x # ys, zs)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   547
  | "\<not> f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, ys, x # zs)"
32314
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   548
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   549
code_pred (expected_modes: (i => bool) => i => bool, (i => bool) => (i * i * o) => bool, (i => bool) => (i * o * i) => bool,
33621
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
   550
  (i => bool) => (o * i * i) => bool, (i => bool) => (i * o * o) => bool) tupled_partition .
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   551
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   552
thm tupled_partition.equation
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   553
33628
ed2111a5c3ed renaming code_pred_intros to code_pred_intro
bulwahn
parents: 33626
diff changeset
   554
lemma [code_pred_intro]:
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   555
  "r a b \<Longrightarrow> tranclp r a b"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   556
  "r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   557
  by auto
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   558
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   559
subsection {* transitive predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   560
33753
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   561
text {* Also look at the tabled transitive closure in the Library *}
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   562
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   563
code_pred (modes: (i => o => bool) => i => i => bool, (i => o => bool) => i => o => bool as forwards_trancl,
33621
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
   564
  (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,
dd564a26fd2f adopted predicate compiler examples to new syntax for modes
bulwahn
parents: 33618
diff changeset
   565
  (o => o => bool) => o => i => bool, (o => o => bool) => o => o => bool) tranclp
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   566
proof -
31580
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31575
diff changeset
   567
  case tranclp
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31575
diff changeset
   568
  from this converse_tranclpE[OF this(1)] show thesis by metis
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   569
qed
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   570
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   571
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   572
code_pred [dseq] tranclp .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   573
code_pred [random_dseq] tranclp .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   574
thm tranclp.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   575
thm tranclp.random_dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   576
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   577
inductive rtrancl' :: "'a => 'a => ('a => 'a => bool) => bool" 
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   578
where
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   579
  "rtrancl' x x r"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   580
| "r x y ==> rtrancl' y z r ==> rtrancl' x z r"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   581
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   582
code_pred [random_dseq] rtrancl' .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   583
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   584
thm rtrancl'.random_dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   585
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   586
inductive rtrancl'' :: "('a * 'a * ('a \<Rightarrow> 'a \<Rightarrow> bool)) \<Rightarrow> bool"  
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   587
where
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   588
  "rtrancl'' (x, x, r)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   589
| "r x y \<Longrightarrow> rtrancl'' (y, z, r) \<Longrightarrow> rtrancl'' (x, z, r)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   590
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   591
code_pred rtrancl'' .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   592
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   593
inductive rtrancl''' :: "('a * ('a * 'a) * ('a * 'a => bool)) => bool" 
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   594
where
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   595
  "rtrancl''' (x, (x, x), r)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   596
| "r (x, y) ==> rtrancl''' (y, (z, z), r) ==> rtrancl''' (x, (z, z), r)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   597
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   598
code_pred rtrancl''' .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   599
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   600
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   601
inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   602
    "succ 0 1"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   603
  | "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   604
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   605
code_pred (modes: i => i => bool, i => o => bool, o => i => bool, o => o => bool) succ .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   606
code_pred [random_dseq] succ .
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   607
thm succ.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   608
thm succ.random_dseq_equation
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   609
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   610
values 10 "{(m, n). succ n m}"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   611
values "{m. succ 0 m}"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   612
values "{m. succ m 0}"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   613
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33477
diff changeset
   614
text {* values command needs mode annotation of the parameter succ
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33477
diff changeset
   615
to disambiguate which mode is to be chosen. *} 
33626
42f69386943a new names for predicate functions in the predicate compiler
bulwahn
parents: 33624
diff changeset
   616
42f69386943a new names for predicate functions in the predicate compiler
bulwahn
parents: 33624
diff changeset
   617
values [mode: i => o => bool] 20 "{n. tranclp succ 10 n}"
42f69386943a new names for predicate functions in the predicate compiler
bulwahn
parents: 33624
diff changeset
   618
values [mode: o => i => bool] 10 "{n. tranclp succ n 10}"
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   619
values 20 "{(n, m). tranclp succ n m}"
33479
428ddcc16e7b added optional mode annotations for parameters in the values command
bulwahn
parents: 33477
diff changeset
   620
33753
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   621
inductive example_graph :: "int => int => bool"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   622
where
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   623
  "example_graph 0 1"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   624
| "example_graph 1 2"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   625
| "example_graph 1 3"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   626
| "example_graph 4 7"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   627
| "example_graph 4 5"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   628
| "example_graph 5 6"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   629
| "example_graph 7 6"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   630
| "example_graph 7 8"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   631
 
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   632
inductive not_reachable_in_example_graph :: "int => int => bool"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   633
where "\<not> (tranclp example_graph x y) ==> not_reachable_in_example_graph x y"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   634
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   635
code_pred (expected_modes: i => i => bool) not_reachable_in_example_graph .
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   636
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   637
thm not_reachable_in_example_graph.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   638
thm tranclp.equation
33753
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   639
value "not_reachable_in_example_graph 0 3"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   640
value "not_reachable_in_example_graph 4 8"
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   641
value "not_reachable_in_example_graph 5 6"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   642
text {* rtrancl compilation is strange! *}
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   643
(*
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   644
value "not_reachable_in_example_graph 0 4"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   645
value "not_reachable_in_example_graph 1 6"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   646
value "not_reachable_in_example_graph 8 4"*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   647
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   648
code_pred [dseq] not_reachable_in_example_graph .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   649
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   650
values [dseq 6] "{x. tranclp example_graph 0 3}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   651
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   652
values [dseq 0] "{x. not_reachable_in_example_graph 0 3}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   653
values [dseq 0] "{x. not_reachable_in_example_graph 0 4}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   654
values [dseq 20] "{x. not_reachable_in_example_graph 0 4}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   655
values [dseq 6] "{x. not_reachable_in_example_graph 0 3}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   656
values [dseq 3] "{x. not_reachable_in_example_graph 4 2}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   657
values [dseq 6] "{x. not_reachable_in_example_graph 4 2}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   658
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   659
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   660
inductive not_reachable_in_example_graph' :: "int => int => bool"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   661
where "\<not> (rtranclp example_graph x y) ==> not_reachable_in_example_graph' x y"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   662
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   663
code_pred not_reachable_in_example_graph' .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   664
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   665
value "not_reachable_in_example_graph' 0 3"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   666
(* value "not_reachable_in_example_graph' 0 5" would not terminate *)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   667
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   668
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   669
(*values [depth_limited 0] "{x. not_reachable_in_example_graph' 0 3}"*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   670
(*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   671
(*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   672
(*values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   673
(*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   674
(*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   675
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   676
code_pred [dseq] not_reachable_in_example_graph' .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   677
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   678
(*thm not_reachable_in_example_graph'.dseq_equation*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   679
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   680
(*values [dseq 0] "{x. not_reachable_in_example_graph' 0 3}"*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   681
(*values [depth_limited 3] "{x. not_reachable_in_example_graph' 0 3}"*) (* fails with undefined *)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   682
(*values [depth_limited 5] "{x. not_reachable_in_example_graph' 0 3}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   683
values [depth_limited 1] "{x. not_reachable_in_example_graph' 0 4}"*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   684
(*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   685
(*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   686
33753
1344e9bb611e changing the proof procedure for parameters; adding a testcase for negation and parameters; adopting print_tac to the latest function print_tac' in the predicate compiler
bulwahn
parents: 33752
diff changeset
   687
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   688
subsection {* IMP *}
32424
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   689
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   690
types
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   691
  var = nat
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   692
  state = "int list"
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   693
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   694
datatype com =
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   695
  Skip |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   696
  Ass var "state => int" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   697
  Seq com com |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   698
  IF "state => bool" com com |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   699
  While "state => bool" com
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   700
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   701
inductive exec :: "com => state => state => bool" where
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   702
"exec Skip s s" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   703
"exec (Ass x e) s (s[x := e(s)])" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   704
"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   705
"b s ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   706
"~b s ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   707
"~b s ==> exec (While b c) s s" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   708
"b s1 ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3"
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   709
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   710
code_pred exec .
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   711
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   712
values "{t. exec
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   713
 (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))))
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   714
 [3,5] t}"
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   715
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   716
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   717
inductive tupled_exec :: "(com \<times> state \<times> state) \<Rightarrow> bool" where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   718
"tupled_exec (Skip, s, s)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   719
"tupled_exec (Ass x e, s, s[x := e(s)])" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   720
"tupled_exec (c1, s1, s2) ==> tupled_exec (c2, s2, s3) ==> tupled_exec (Seq c1 c2, s1, s3)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   721
"b s ==> tupled_exec (c1, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   722
"~b s ==> tupled_exec (c2, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   723
"~b s ==> tupled_exec (While b c, s, s)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   724
"b s1 ==> tupled_exec (c, s1, s2) ==> tupled_exec (While b c, s2, s3) ==> tupled_exec (While b c, s1, s3)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   725
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   726
code_pred tupled_exec .
32424
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   727
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   728
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)}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   729
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   730
subsection {* CCS *}
32408
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   731
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   732
text{* This example formalizes finite CCS processes without communication or
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   733
recursion. For simplicity, labels are natural numbers. *}
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   734
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   735
datatype proc = nil | pre nat proc | or proc proc | par proc proc
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   736
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   737
inductive step :: "proc \<Rightarrow> nat \<Rightarrow> proc \<Rightarrow> bool" where
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   738
"step (pre n p) n p" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   739
"step p1 a q \<Longrightarrow> step (or p1 p2) a q" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   740
"step p2 a q \<Longrightarrow> step (or p1 p2) a q" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   741
"step p1 a q \<Longrightarrow> step (par p1 p2) a (par q p2)" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   742
"step p2 a q \<Longrightarrow> step (par p1 p2) a (par p1 q)"
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   743
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   744
code_pred step .
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   745
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   746
inductive steps where
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   747
"steps p [] p" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   748
"step p a q \<Longrightarrow> steps q as r \<Longrightarrow> steps p (a#as) r"
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   749
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   750
code_pred steps .
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   751
33477
1272cfc7b910 added further example of the values command
bulwahn
parents: 33475
diff changeset
   752
values 3 
1272cfc7b910 added further example of the values command
bulwahn
parents: 33475
diff changeset
   753
 "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
1272cfc7b910 added further example of the values command
bulwahn
parents: 33475
diff changeset
   754
32408
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   755
values 5
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   756
 "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   757
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   758
values 3 "{(a,q). step (par nil nil) a q}"
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   759
32408
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   760
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   761
inductive tupled_step :: "(proc \<times> nat \<times> proc) \<Rightarrow> bool"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   762
where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   763
"tupled_step (pre n p, n, p)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   764
"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   765
"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   766
"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par q p2)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   767
"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par p1 q)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   768
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   769
code_pred tupled_step .
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   770
thm tupled_step.equation
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   771
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   772
subsection {* divmod *}
32579
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   773
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   774
inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   775
    "k < l \<Longrightarrow> divmod_rel k l 0 k"
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   776
  | "k \<ge> l \<Longrightarrow> divmod_rel (k - l) l q r \<Longrightarrow> divmod_rel k l (Suc q) r"
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   777
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   778
code_pred divmod_rel ..
33626
42f69386943a new names for predicate functions in the predicate compiler
bulwahn
parents: 33624
diff changeset
   779
thm divmod_rel.equation
42f69386943a new names for predicate functions in the predicate compiler
bulwahn
parents: 33624
diff changeset
   780
value [code] "Predicate.the (divmod_rel_i_i_o_o 1705 42)"
32579
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   781
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   782
subsection {* Minimum *}
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   783
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   784
definition Min
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   785
where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   786
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   787
code_pred [inductify] Min .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   788
thm Min.equation
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   789
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   790
subsection {* Lexicographic order *}
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   791
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   792
declare lexord_def[code_pred_def]
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   793
code_pred [inductify] lexord .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   794
code_pred [random_dseq inductify] lexord .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   795
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   796
thm lexord.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   797
thm lexord.random_dseq_equation
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   798
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   799
inductive less_than_nat :: "nat * nat => bool"
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   800
where
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   801
  "less_than_nat (0, x)"
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   802
| "less_than_nat (x, y) ==> less_than_nat (Suc x, Suc y)"
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   803
 
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   804
code_pred less_than_nat .
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   805
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   806
code_pred [dseq] less_than_nat .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   807
code_pred [random_dseq] less_than_nat .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   808
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   809
inductive test_lexord :: "nat list * nat list => bool"
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   810
where
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   811
  "lexord less_than_nat (xs, ys) ==> test_lexord (xs, ys)"
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   812
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   813
code_pred test_lexord .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   814
code_pred [dseq] test_lexord .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   815
code_pred [random_dseq] test_lexord .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   816
thm test_lexord.dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   817
thm test_lexord.random_dseq_equation
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   818
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   819
values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   820
(*values [depth_limited 5] "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"*)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   821
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   822
declare list.size(3,4)[code_pred_def]
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   823
lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   824
(*
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   825
code_pred [inductify] lexn .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   826
thm lexn.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   827
*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   828
(*
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   829
code_pred [random_dseq inductify] lexn .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   830
thm lexn.random_dseq_equation
33139
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
   831
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   832
values [random_dseq 4, 4, 6] 100 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   833
*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   834
inductive has_length
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   835
where
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   836
  "has_length [] 0"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   837
| "has_length xs i ==> has_length (x # xs) (Suc i)" 
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   838
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   839
lemma has_length:
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   840
  "has_length xs n = (length xs = n)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   841
proof (rule iffI)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   842
  assume "has_length xs n"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   843
  from this show "length xs = n"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   844
    by (rule has_length.induct) auto
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   845
next
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   846
  assume "length xs = n"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   847
  from this show "has_length xs n"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   848
    by (induct xs arbitrary: n) (auto intro: has_length.intros)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   849
qed
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   850
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   851
lemma lexn_intros [code_pred_intro]:
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   852
  "has_length xs i ==> has_length ys i ==> r (x, y) ==> lexn r (Suc i) (x # xs, y # ys)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   853
  "lexn r i (xs, ys) ==> lexn r (Suc i) (x # xs, x # ys)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   854
proof -
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   855
  assume "has_length xs i" "has_length ys i" "r (x, y)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   856
  from this has_length show "lexn r (Suc i) (x # xs, y # ys)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   857
    unfolding lexn_conv Collect_def mem_def
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   858
    by fastsimp
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   859
next
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   860
  assume "lexn r i (xs, ys)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   861
  thm lexn_conv
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   862
  from this show "lexn r (Suc i) (x#xs, x#ys)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   863
    unfolding Collect_def mem_def lexn_conv
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   864
    apply auto
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   865
    apply (rule_tac x="x # xys" in exI)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   866
    by auto
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   867
qed
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   868
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   869
code_pred [random_dseq inductify] lexn
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   870
proof -
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   871
  fix n xs ys
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   872
  assume 1: "lexn r n (xs, ys)"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   873
  assume 2: "\<And>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"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   874
  assume 3: "\<And>i x xs' ys'. r = r ==> n = Suc i ==> (xs, ys) = (x # xs', x # ys') ==> lexn r i (xs', ys') ==> thesis"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   875
  from 1 2 3 show thesis   
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   876
    unfolding lexn_conv Collect_def mem_def
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   877
    apply (auto simp add: has_length)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   878
    apply (case_tac xys)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   879
    apply auto
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   880
    apply fastsimp
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   881
    apply fastsimp done
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   882
qed
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   883
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   884
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   885
values [random_dseq 1, 2, 5] 10 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   886
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   887
33375
fd3e861f8d31 renamed rpred to random
bulwahn
parents: 33329
diff changeset
   888
code_pred [inductify] lenlex .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   889
thm lenlex.equation
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   890
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   891
code_pred [random_dseq inductify] lenlex .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   892
thm lenlex.random_dseq_equation
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   893
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   894
values [random_dseq 4, 2, 4] 100 "{(xs, ys::int list). lenlex (%(x, y). x <= y) (xs, ys)}"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   895
thm lists.intros
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   896
(*
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   897
code_pred [inductify] lists .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   898
*)
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   899
(*thm lists.equation*)
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   900
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   901
subsection {* AVL Tree *}
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   902
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   903
datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   904
fun height :: "'a tree => nat" where
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   905
"height ET = 0"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   906
| "height (MKT x l r h) = max (height l) (height r) + 1"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   907
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   908
consts avl :: "'a tree => bool"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   909
primrec
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   910
  "avl ET = True"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   911
  "avl (MKT x l r h) = ((height l = height r \<or> height l = 1 + height r \<or> height r = 1+height l) \<and> 
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   912
  h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   913
(*
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   914
code_pred [inductify] avl .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   915
thm avl.equation*)
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   916
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   917
code_pred [random_dseq inductify] avl .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   918
thm avl.random_dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   919
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   920
values [random_dseq 2, 1, 7] 5 "{t:: int tree. avl t}"
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   921
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   922
fun set_of
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   923
where
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   924
"set_of ET = {}"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   925
| "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   926
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33120
diff changeset
   927
fun is_ord :: "nat tree => bool"
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   928
where
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   929
"is_ord ET = True"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   930
| "is_ord (MKT n l r h) =
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   931
 ((\<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)"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   932
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   933
code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] set_of .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   934
thm set_of.equation
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   935
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   936
code_pred (expected_modes: i => bool) [inductify] is_ord .
33480
dcfe9100e0a6 disabled upt example because of a problem due to overloaded constants with the predicate compiler
bulwahn
parents: 33479
diff changeset
   937
thm is_ord_aux.equation
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   938
thm is_ord.equation
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   939
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
   940
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   941
subsection {* Definitions about Relations *}
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   942
term "converse"
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   943
code_pred (modes:
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   944
  (i * i => bool) => i * i => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   945
  (i * o => bool) => o * i => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   946
  (i * o => bool) => i * i => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   947
  (o * i => bool) => i * o => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   948
  (o * i => bool) => i * i => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   949
  (o * o => bool) => o * o => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   950
  (o * o => bool) => i * o => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   951
  (o * o => bool) => o * i => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   952
  (o * o => bool) => i * i => bool) [inductify] converse .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   953
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   954
thm converse.equation
33145
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   955
code_pred [inductify] rel_comp .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   956
thm rel_comp.equation
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   957
code_pred [inductify] Image .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   958
thm Image.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   959
declare singleton_iff[code_pred_inline]
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   960
declare Id_on_def[unfolded Bex_def UNION_def singleton_iff, code_pred_def]
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   961
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   962
code_pred (expected_modes:
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   963
  (o => bool) => o => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   964
  (o => bool) => i * o => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   965
  (o => bool) => o * i => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   966
  (o => bool) => i => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   967
  (i => bool) => i * o => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   968
  (i => bool) => o * i => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   969
  (i => bool) => i => bool) [inductify] Id_on .
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   970
thm Id_on.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   971
thm Domain_def
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   972
code_pred (modes:
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   973
  (o * o => bool) => o => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   974
  (o * o => bool) => i => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   975
  (i * o => bool) => i => bool) [inductify] Domain .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   976
thm Domain.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   977
code_pred (modes:
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   978
  (o * o => bool) => o => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   979
  (o * o => bool) => i => bool,
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   980
  (o * i => bool) => i => bool) [inductify] Range .
33253
d9ca0c3bf680 changes to example file
bulwahn
parents: 33250
diff changeset
   981
thm Range.equation
33145
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   982
code_pred [inductify] Field .
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
   983
thm Field.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   984
(*thm refl_on_def
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   985
code_pred [inductify] refl_on .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
   986
thm refl_on.equation*)
33145
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   987
code_pred [inductify] total_on .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   988
thm total_on.equation
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   989
code_pred [inductify] antisym .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   990
thm antisym.equation
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   991
code_pred [inductify] trans .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   992
thm trans.equation
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   993
code_pred [inductify] single_valued .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   994
thm single_valued.equation
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   995
code_pred [inductify] inv_image .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   996
thm inv_image.equation
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   997
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   998
subsection {* Inverting list functions *}
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
   999
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1000
(*code_pred [inductify] length .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1001
code_pred [random inductify] length .
33114
4785ef554dcc added further examples; added mode to code_pred command; tuned; some temporary things in Predicate_Compile_ex
bulwahn
parents: 33113
diff changeset
  1002
thm size_listP.equation
33375
fd3e861f8d31 renamed rpred to random
bulwahn
parents: 33329
diff changeset
  1003
thm size_listP.random_equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1004
*)
33405
5c1928d5db38 adding testcases for code inlining, sets and numerals to the example files of the predicate compiler and its quickcheck prototype; disabling length example
bulwahn
parents: 33375
diff changeset
  1005
(*values [random] 1 "{xs. size_listP (xs::nat list) (5::nat)}"*)
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
  1006
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
  1007
code_pred (expected_modes: i => o => bool, o => i => bool, i => i => bool) [inductify] List.concat .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1008
thm concatP.equation
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1009
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1010
values "{ys. concatP [[1, 2], [3, (4::int)]] ys}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1011
values "{ys. concatP [[1, 2], [3]] [1, 2, (3::nat)]}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1012
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1013
code_pred [dseq inductify] List.concat .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1014
thm concatP.dseq_equation
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1015
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1016
values [dseq 3] 3
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1017
  "{xs. concatP xs ([0] :: nat list)}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1018
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1019
values [dseq 5] 3
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1020
  "{xs. concatP xs ([1] :: int list)}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1021
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1022
values [dseq 5] 3
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1023
  "{xs. concatP xs ([1] :: nat list)}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1024
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1025
values [dseq 5] 3
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1026
  "{xs. concatP xs [(1::int), 2]}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1027
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
  1028
code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] hd .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1029
thm hdP.equation
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1030
values "{x. hdP [1, 2, (3::int)] x}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1031
values "{(xs, x). hdP [1, 2, (3::int)] 1}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1032
 
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
  1033
code_pred (expected_modes: i => o => bool, i => i => bool) [inductify] tl .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1034
thm tlP.equation
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1035
values "{x. tlP [1, 2, (3::nat)] x}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1036
values "{x. tlP [1, 2, (3::int)] [3]}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1037
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1038
code_pred [inductify] last .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1039
thm lastP.equation
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1040
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1041
code_pred [inductify] butlast .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1042
thm butlastP.equation
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1043
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1044
code_pred [inductify] take .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1045
thm takeP.equation
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1046
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1047
code_pred [inductify] drop .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1048
thm dropP.equation
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1049
code_pred [inductify] zip .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1050
thm zipP.equation
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1051
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1052
code_pred [inductify] upt .
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1053
code_pred [inductify] remdups .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1054
thm remdupsP.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1055
code_pred [dseq inductify] remdups .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1056
values [dseq 4] 5 "{xs. remdupsP xs [1, (2::int)]}"
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1057
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1058
code_pred [inductify] remove1 .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1059
thm remove1P.equation
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1060
values "{xs. remove1P 1 xs [2, (3::int)]}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1061
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1062
code_pred [inductify] removeAll .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1063
thm removeAllP.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1064
code_pred [dseq inductify] removeAll .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1065
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1066
values [dseq 4] 10 "{xs. removeAllP 1 xs [(2::nat)]}"
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1067
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1068
code_pred [inductify] distinct .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1069
thm distinct.equation
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1070
code_pred [inductify] replicate .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1071
thm replicateP.equation
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1072
values 5 "{(n, xs). replicateP n (0::int) xs}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1073
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1074
code_pred [inductify] splice .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1075
thm splice.simps
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1076
thm spliceP.equation
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1077
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1078
values "{xs. spliceP xs [1, 2, 3] [1, 1, 1, 2, 1, (3::nat)]}"
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1079
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1080
code_pred [inductify] List.rev .
33129
3085da75ed54 changed import_intros to handle parameters differently; changed handling of higher-order function compilation; reverted MicroJava change; tuned
bulwahn
parents: 33128
diff changeset
  1081
code_pred [inductify] map .
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1082
code_pred [inductify] foldr .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1083
code_pred [inductify] foldl .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
  1084
code_pred [inductify] filter .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1085
code_pred [random_dseq inductify] filter .
33141
89b23fad5e02 modifying the depth-limited compilation to be sound, but now throws an error undefined in case of hitting the depth limit in an negative context; cleaning up the examples
bulwahn
parents: 33140
diff changeset
  1086
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
  1087
subsection {* Context Free Grammar *}
32669
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1088
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1089
datatype alphabet = a | b
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1090
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1091
inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1092
  "[] \<in> S\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1093
| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1094
| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1095
| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1096
| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1097
| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1098
33140
83951822bfd0 cleaning the signature of the predicate compiler core; renaming signature and structures to uniform long names
bulwahn
parents: 33139
diff changeset
  1099
code_pred [inductify] S\<^isub>1p .
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1100
code_pred [random_dseq inductify] S\<^isub>1p .
32669
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1101
thm S\<^isub>1p.equation
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1102
thm S\<^isub>1p.random_dseq_equation
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1103
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1104
values [random_dseq 5, 5, 5] 5 "{x. S\<^isub>1p x}"
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1105
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1106
inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1107
  "[] \<in> S\<^isub>2"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1108
| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1109
| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1110
| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1111
| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1112
| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1113
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1114
code_pred [random_dseq inductify] S\<^isub>2p .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1115
thm S\<^isub>2p.random_dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1116
thm A\<^isub>2p.random_dseq_equation
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1117
thm B\<^isub>2p.random_dseq_equation
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1118
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1119
values [random_dseq 5, 5, 5] 10 "{x. S\<^isub>2p x}"
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1120
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1121
inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1122
  "[] \<in> S\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1123
| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1124
| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1125
| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1126
| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1127
| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1128
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1129
code_pred [inductify] S\<^isub>3p .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1130
thm S\<^isub>3p.equation
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1131
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1132
values 10 "{x. S\<^isub>3p x}"
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
  1133
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1134
inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1135
  "[] \<in> S\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1136
| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1137
| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1138
| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1139
| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1140
| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1141
| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1142
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
  1143
code_pred (expected_modes: o => bool, i => bool) S\<^isub>4p .
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
  1144
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
  1145
subsection {* Lambda *}
32669
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1146
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1147
datatype type =
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1148
    Atom nat
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1149
  | Fun type type    (infixr "\<Rightarrow>" 200)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1150
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1151
datatype dB =
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1152
    Var nat
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1153
  | App dB dB (infixl "\<degree>" 200)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1154
  | Abs type dB
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1155
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1156
primrec
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1157
  nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1158
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1159
  "[]\<langle>i\<rangle> = None"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1160
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1161
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1162
inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1163
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1164
  "nth_el' (x # xs) 0 x"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1165
| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
33128
1f990689349f further cleaning up
bulwahn
parents: 33126
diff changeset
  1166
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1167
inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile> _ : _" [50, 50, 50] 50)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1168
  where
33128
1f990689349f further cleaning up
bulwahn
parents: 33126
diff changeset
  1169
    Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1170
  | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
  1171
  | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1172
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1173
primrec
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1174
  lift :: "[dB, nat] => dB"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1175
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1176
    "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1177
  | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1178
  | "lift (Abs T s) k = Abs T (lift s (k + 1))"
32669
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
  1179
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1180
primrec
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1181
  subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1182
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1183
    subst_Var: "(Var i)[s/k] =
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1184
      (if k < i then Var (i - 1) else if i = k then s else Var i)"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1185
  | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1186
  | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1187
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1188
inductive beta :: "[dB, dB] => bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1189
  where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1190
    beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1191
  | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1192
  | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1193
  | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
  1194
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
  1195
code_pred (expected_modes: i => i => o => bool, i => i => i => bool) typing .
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
  1196
thm typing.equation
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
  1197
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1198
code_pred (modes: i => i => bool,  i => o => bool as reduce') beta .
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
  1199
thm beta.equation
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
  1200
33623
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
  1201
values "{x. App (Abs (Atom 0) (Var 0)) (Var 1) \<rightarrow>\<^sub>\<beta> x}"
4ec42d38224f changed modes to expected_modes; added UNION to code_pred_inlining; fixed some examples; tuned
bulwahn
parents: 33621
diff changeset
  1202
33624
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1203
definition "reduce t = Predicate.the (reduce' t)"
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1204
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1205
value "reduce (App (Abs (Atom 0) (Var 0)) (Var 1))"
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1206
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1207
code_pred [dseq] typing .
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1208
code_pred [random_dseq] typing .
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1209
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1210
values [random_dseq 1,1,5] 10 "{(\<Gamma>, t, T). \<Gamma> \<turnstile> t : T}"
33618
d8359a16e0c5 adding more tests for the values command; adding some forbidden constants to inductify
bulwahn
parents: 33480
diff changeset
  1211
33624
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1212
subsection {* A minimal example of yet another semantics *}
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1213
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1214
text {* thanks to Elke Salecker *}
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1215
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1216
types
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1217
  vname = nat
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1218
  vvalue = int
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1219
  var_assign = "vname \<Rightarrow> vvalue"  --"variable assignment"
33624
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1220
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1221
datatype ir_expr = 
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1222
  IrConst vvalue
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1223
| ObjAddr vname
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1224
| Add ir_expr ir_expr
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1225
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1226
datatype val =
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1227
  IntVal  vvalue
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1228
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1229
record  configuration =
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1230
  Env :: var_assign
33624
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1231
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1232
inductive eval_var ::
34948
2d5f2a9f7601 refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents: 33753
diff changeset
  1233
  "ir_expr \<Rightarrow> configuration \<Rightarrow> val \<Rightarrow> bool"
33624
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1234
where
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1235
  irconst: "eval_var (IrConst i) conf (IntVal i)"
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1236
| objaddr: "\<lbrakk> Env conf n = i \<rbrakk> \<Longrightarrow> eval_var (ObjAddr n) conf (IntVal i)"
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1237
| 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))"
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1238
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1239
33752
9aa8e961f850 adopting proposed_modes; adding a new dimension of complexity for nicer error messages; tuned
bulwahn
parents: 33628
diff changeset
  1240
code_pred eval_var .
33624
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1241
thm eval_var.equation
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1242
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1243
values "{val. eval_var (Add (IrConst 1) (IrConst 2)) (| Env = (\<lambda>x. 0)|) val}"
a68a391a1451 added another example to the predicate compiler
bulwahn
parents: 33623
diff changeset
  1244
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
  1245
end