src/HOL/ex/Predicate_Compile_ex.thy
author bulwahn
Wed, 28 Oct 2009 12:29:01 +0100
changeset 33327 9d03957622a2
parent 33326 7d0288d90535
child 33329 b129e4c476d6
permissions -rw-r--r--
improving mode parsing in the predicate compiler
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
33327
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
     2
imports "../Main" 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
33327
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
     9
code_pred (mode: X ! []) False' .
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    10
code_pred [depth_limited] False' .
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
    11
code_pred [rpred] False' .
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    12
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    13
inductive EmptySet :: "'a \<Rightarrow> bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    14
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    15
code_pred (mode: [], [1]) EmptySet .
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    16
33253
d9ca0c3bf680 changes to example file
bulwahn
parents: 33250
diff changeset
    17
definition EmptySet' :: "'a \<Rightarrow> bool"
d9ca0c3bf680 changes to example file
bulwahn
parents: 33250
diff changeset
    18
where "EmptySet' = {}"
d9ca0c3bf680 changes to example file
bulwahn
parents: 33250
diff changeset
    19
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    20
code_pred (mode: [], [1]) [inductify] EmptySet' .
33253
d9ca0c3bf680 changes to example file
bulwahn
parents: 33250
diff changeset
    21
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    22
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
    23
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    24
code_pred (mode: [], [1], [2], [1, 2]) EmptyRel .
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    25
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    26
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
    27
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
    28
33327
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
    29
code_pred
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
    30
(*  (mode: [] => [], [] => [1], [] => [2], [] => [1, 2],
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
    31
         [1] => [], [1] => [1], [1] => [2], [1] => [1, 2],
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
    32
         [2] => [], [2] => [1], [2] => [2], [2] => [1, 2],
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
    33
         [1, 2] => [], [1, 2] => [1], [1, 2] => [2], [1, 2] => [1, 2])*)
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
    34
  EmptyClosure .
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
    35
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    36
thm EmptyClosure.equation
33258
f47ca46d2187 disabled test case for predicate compiler due to an problem in the inductive package
bulwahn
parents: 33255
diff changeset
    37
(* TODO: inductive package is broken!
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    38
inductive False'' :: "bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    39
where
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    40
  "False \<Longrightarrow> False''"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    41
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    42
code_pred (mode: []) False'' .
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    43
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
    44
inductive EmptySet'' :: "'a \<Rightarrow> bool"
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    45
where
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
    46
  "False \<Longrightarrow> EmptySet'' x"
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    47
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
    48
code_pred (mode: [1]) EmptySet'' .
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
    49
code_pred (mode: [], [1]) [inductify] EmptySet'' .
33258
f47ca46d2187 disabled test case for predicate compiler due to an problem in the inductive package
bulwahn
parents: 33255
diff changeset
    50
*)
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    51
inductive True' :: "bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    52
where
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    53
  "True \<Longrightarrow> True'"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    54
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    55
code_pred (mode: []) True' .
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    56
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    57
consts a' :: 'a
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    58
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    59
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
    60
where
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    61
"Fact a' a'"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    62
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    63
code_pred (mode: [], [1], [2], [1, 2]) Fact .
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    64
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    65
inductive zerozero :: "nat * nat => bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    66
where
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    67
  "zerozero (0, 0)"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    68
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
    69
code_pred zerozero .
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    70
code_pred [rpred] zerozero .
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    71
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    72
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
    73
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    74
datatype char = C | D | E | F | G
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    75
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    76
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
    77
where
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    78
  "(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
    79
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    80
code_pred (mode: [1]) 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
    81
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
    82
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    83
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
    84
where
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    85
  "(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
    86
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    87
lemma [code_pred_intros]:
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    88
  "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
    89
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
    90
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    91
lemma [code_pred_intros]:
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    92
  "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
    93
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
    94
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    95
code_pred (mode: [], [1]) 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
    96
proof -
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
    97
  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
    98
  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
    99
  proof
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   100
    fix x
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   101
    assume x: "a1 = x"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   102
    assume "x = D \<or> x = E"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   103
    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
   104
    proof
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   105
      assume "x = D" from this x is_D_or_E(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
   106
    next
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   107
      assume "x = E" from this x is_D_or_E(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
   108
    qed
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   109
  qed
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   110
qed
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   111
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   112
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
   113
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   114
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
   115
where
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   116
  "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
   117
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   118
lemma [code_pred_intros]:
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   119
  "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
   120
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
   121
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   122
lemma [code_pred_intros]:
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   123
  "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
   124
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
   125
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   126
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
   127
where
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   128
  "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
   129
| "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
   130
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   131
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
   132
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   133
code_pred 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
   134
proof -
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   135
  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
   136
  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
   137
  proof
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   138
    fix x
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   139
    assume x: "a1 = x"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   140
    assume "x = F \<or> x = G"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   141
    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
   142
    proof
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   143
      assume "x = F"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   144
      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
   145
    next
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   146
      assume "x = G"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   147
      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
   148
    qed
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   149
  qed
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   150
qed
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   151
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   152
subsection {* Preprocessor Inlining  *}
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   153
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   154
definition "equals == (op =)"
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   155
 
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   156
inductive zerozero' where
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   157
  "equals (x, y) (0, 0) ==> zerozero' (x, y)"
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   158
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   159
code_pred (mode: [1]) zerozero' .
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   160
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   161
lemma zerozero'_eq: "zerozero' == zerozero"
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   162
proof -
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   163
  have "zerozero' = zerozero"
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   164
    apply (auto simp add: mem_def)
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   165
    apply (cases rule: zerozero'.cases)
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   166
    apply (auto simp add: equals_def intro: zerozero.intros)
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   167
    apply (cases rule: zerozero.cases)
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   168
    apply (auto simp add: equals_def intro: zerozero'.intros)
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   169
    done
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   170
  from this show "zerozero' == zerozero" by auto
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   171
qed
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   172
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   173
declare zerozero'_eq [code_pred_inline]
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   174
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   175
definition "zerozero'' x == zerozero' x"
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   176
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   177
text {* if preprocessing fails, zerozero'' will not have all modes. *}
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   178
ML {* Predicate_Compile_Inline_Defs.get @{context} *}
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   179
(* TODO: *)
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   180
code_pred (mode: [1]) [inductify, show_intermediate_results] zerozero'' .
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   181
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   182
subsection {* even predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   183
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   184
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
   185
    "even 0"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   186
  | "even n \<Longrightarrow> odd (Suc n)"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   187
  | "odd n \<Longrightarrow> even (Suc n)"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   188
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
   189
code_pred (mode: [], [1]) even .
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
   190
code_pred [depth_limited] even .
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
   191
code_pred [rpred] even .
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 33136
diff changeset
   192
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
   193
thm odd.equation
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   194
thm even.equation
33136
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   195
thm odd.depth_limited_equation
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   196
thm even.depth_limited_equation
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   197
thm even.rpred_equation
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   198
thm odd.rpred_equation
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   199
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   200
values "{x. even 2}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   201
values "{x. odd 2}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   202
values 10 "{n. even n}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   203
values 10 "{n. odd n}"
33136
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   204
values [depth_limit = 2] "{x. even 6}"
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   205
values [depth_limit = 7] "{x. even 6}"
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   206
values [depth_limit = 2] "{x. odd 7}"
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   207
values [depth_limit = 8] "{x. odd 7}"
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   208
values [depth_limit = 7] 10 "{n. even n}"
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
   209
33136
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   210
definition odd' where "odd' x == \<not> even x"
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   211
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   212
code_pred [inductify] odd' .
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   213
code_pred [inductify, depth_limited] odd' .
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
   214
code_pred [inductify, rpred] odd' .
33136
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   215
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   216
thm odd'.depth_limited_equation
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   217
values [depth_limit = 2] "{x. odd' 7}"
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   218
values [depth_limit = 9] "{x. odd' 7}"
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   219
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   220
inductive is_even :: "nat \<Rightarrow> bool"
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   221
where
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   222
  "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
   223
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   224
code_pred is_even .
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   225
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   226
subsection {* append predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   227
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   228
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   229
    "append [] xs xs"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   230
  | "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
   231
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   232
code_pred (mode: [1, 2], [3], [2, 3], [1, 3], [1, 2, 3]) append .
33136
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   233
code_pred [depth_limited] append .
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
   234
code_pred [rpred] append .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   235
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   236
thm append.equation
33136
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   237
thm append.depth_limited_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
   238
thm append.rpred_equation
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   239
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   240
values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   241
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
   242
values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0, 5]}"
33136
74d51fb3be2e commented out the random generator compilation in the example file
bulwahn
parents: 33129
diff changeset
   243
values [depth_limit = 3] "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
   244
values [random] 15 "{(ys, zs). append [1::nat, 2] ys zs}"
31195
12741f23527d added example on ML level
haftmann
parents: 31129
diff changeset
   245
33111
db5af7b86a2f developing an executable the operator
bulwahn
parents: 33108
diff changeset
   246
value [code] "Predicate.the (append_1_2 [0::int, 1, 2] [3, 4, 5])"
db5af7b86a2f developing an executable the operator
bulwahn
parents: 33108
diff changeset
   247
value [code] "Predicate.the (append_3 ([]::int list))"
db5af7b86a2f developing an executable the operator
bulwahn
parents: 33108
diff changeset
   248
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   249
text {* tricky case with alternative rules *}
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   250
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
   251
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
   252
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
   253
  "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
   254
| "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
   255
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
   256
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
   257
  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
   258
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
   259
lemmas [code_pred_intros] = append2_Nil append2.intros(2)
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
   260
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
   261
code_pred 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
   262
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
   263
  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
   264
  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
   265
  proof
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   266
    fix xs
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   267
    assume "a1 = []" "a2 = xs" "a3 = xs"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   268
    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
   269
  next
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   270
    fix xs ys zs x
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   271
    assume "a1 = x # xs" "a2 = ys" "a3 = x # zs" "append2 xs ys zs"
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   272
    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
   273
  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
   274
qed
33128
1f990689349f further cleaning up
bulwahn
parents: 33126
diff changeset
   275
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   276
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
   277
where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   278
  "tupled_append ([], xs, xs)"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   279
| "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
   280
33327
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
   281
code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) tupled_append .
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
   282
code_pred [rpred] tupled_append .
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   283
thm tupled_append.equation
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   284
(*
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
   285
TODO: values with tupled modes
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   286
values "{xs. tupled_append ([1,2,3], [4,5], xs)}"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   287
*)
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
   288
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
   289
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
   290
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
   291
"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
   292
| "[| 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
   293
 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
   294
33327
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
   295
code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) 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
   296
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
   297
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
   298
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
   299
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
   300
  "tupled_append'' ([], xs, xs)"
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   301
| "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
   302
33327
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
   303
code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) [inductify] 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
   304
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
   305
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
   306
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
   307
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
   308
  "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
   309
| "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
   310
33327
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
   311
code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) [inductify] 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
   312
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
   313
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   314
subsection {* map_ofP predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   315
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   316
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
   317
where
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   318
  "map_ofP ((a, b)#xs) a b"
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   319
| "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
   320
33120
ca77d8c34ce2 cleaned up
bulwahn
parents: 33119
diff changeset
   321
code_pred (mode: [1], [1, 2], [1, 2, 3], [1, 3]) map_ofP .
33113
0f6e30b87cf1 processing of tuples in introduction rules
bulwahn
parents: 33112
diff changeset
   322
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
   323
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   324
subsection {* filter predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   325
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   326
inductive filter1
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   327
for P
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   328
where
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   329
  "filter1 P [] []"
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   330
| "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
   331
| "\<not> P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys"
33327
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
   332
ML {* Scan.optional *}
9d03957622a2 improving mode parsing in the predicate compiler
bulwahn
parents: 33326
diff changeset
   333
code_pred ( mode : [1], [1, 2]) filter1 .
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   334
code_pred [depth_limited] filter1 .
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   335
code_pred [rpred] filter1 .
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   336
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   337
thm filter1.equation
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   338
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   339
inductive filter2
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   340
where
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   341
  "filter2 P [] []"
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   342
| "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
   343
| "\<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
   344
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   345
code_pred (mode: [1, 2, 3], [1, 2]) filter2 .
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   346
code_pred [depth_limited] filter2 .
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   347
code_pred [rpred] filter2 .
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   348
thm filter2.equation
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   349
thm filter2.rpred_equation
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   350
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   351
inductive filter3
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   352
for P
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   353
where
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   354
  "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
   355
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   356
code_pred filter3 .
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   357
code_pred [depth_limited] filter3 .
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   358
thm filter3.depth_limited_equation
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   359
33147
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   360
inductive filter4
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   361
where
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   362
  "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
   363
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   364
code_pred filter4 .
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   365
code_pred [depth_limited] filter4 .
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   366
code_pred [rpred] filter4 .
180dc60bd88c improving the compilation with higher-order arguments in the predicate compiler
bulwahn
parents: 33146
diff changeset
   367
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   368
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
   369
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   370
inductive rev where
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   371
    "rev [] []"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   372
  | "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   373
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   374
code_pred (mode: [1], [2], [1, 2]) rev .
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   375
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   376
thm rev.equation
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   377
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   378
values "{xs. rev [0, 1, 2, 3::nat] xs}"
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   379
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   380
inductive tupled_rev where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   381
  "tupled_rev ([], [])"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   382
| "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
   383
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   384
code_pred tupled_rev .
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   385
thm tupled_rev.equation
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   386
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   387
subsection {* partition predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   388
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   389
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
   390
  for f where
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   391
    "partition f [] [] []"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
   392
  | "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
   393
  | "\<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
   394
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   395
code_pred (mode: [1], [2, 3], [1, 2], [1, 3], [1, 2, 3]) partition .
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
   396
code_pred [depth_limited] partition .
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
   397
code_pred [rpred] partition .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   398
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   399
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
   400
  [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
   401
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
   402
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
   403
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   404
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
   405
  for f where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   406
   "tupled_partition f ([], [], [])"
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   407
  | "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
   408
  | "\<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
   409
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   410
code_pred tupled_partition .
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   411
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   412
thm tupled_partition.equation
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   413
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31514
diff changeset
   414
lemma [code_pred_intros]:
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   415
  "r a b \<Longrightarrow> tranclp r a b"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   416
  "r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   417
  by auto
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   418
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   419
subsection {* transitive predicate *}
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   420
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   421
code_pred tranclp
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   422
proof -
31580
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31575
diff changeset
   423
  case tranclp
1c143f6a2226 added cases to code_pred command
bulwahn
parents: 31575
diff changeset
   424
  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
   425
qed
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   426
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
   427
code_pred [depth_limited] tranclp .
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
   428
code_pred [rpred] tranclp .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
   429
thm tranclp.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
   430
thm tranclp.rpred_equation
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   431
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   432
inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   433
    "succ 0 1"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   434
  | "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   435
31550
b63d253ed9e2 code_pred command now also requires proofs for dependent predicates; changed handling of parameters in introrules of executable function
bulwahn
parents: 31514
diff changeset
   436
code_pred succ .
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
   437
code_pred [rpred] succ .
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   438
thm succ.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
   439
thm succ.rpred_equation
32340
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   440
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   441
values 10 "{(m, n). succ n m}"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   442
values "{m. succ 0 m}"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   443
values "{m. succ m 0}"
b4632820e74c cleaned up
haftmann
parents: 31879
diff changeset
   444
32314
66bbad0bfef9 changed resolving depending predicates and fetching in the predicate compiler
bulwahn
parents: 32310
diff changeset
   445
(* FIXME: why does this not terminate? -- value chooses mode [] --> [1] and then starts enumerating all successors *)
32355
806d2df4d79d properly merged
haftmann
parents: 32351
diff changeset
   446
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
   447
(*
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   448
values 20 "{n. tranclp succ 10 n}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   449
values "{n. tranclp succ n 10}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   450
values 20 "{(n, m). tranclp succ n m}"
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
   451
*)
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   452
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   453
subsection {* IMP *}
32424
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   454
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   455
types
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   456
  var = nat
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   457
  state = "int list"
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   458
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   459
datatype com =
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   460
  Skip |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   461
  Ass var "state => int" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   462
  Seq com com |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   463
  IF "state => bool" com com |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   464
  While "state => bool" com
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   465
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   466
inductive exec :: "com => state => state => bool" where
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   467
"exec Skip s s" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   468
"exec (Ass x e) s (s[x := e(s)])" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   469
"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   470
"b s ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   471
"~b s ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   472
"~b s ==> exec (While b c) s s" |
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   473
"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
   474
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   475
code_pred exec .
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   476
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   477
values "{t. exec
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   478
 (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
   479
 [3,5] t}"
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   480
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   481
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
   482
"tupled_exec (Skip, s, s)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   483
"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
   484
"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
   485
"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
   486
"~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
   487
"~b s ==> tupled_exec (While b c, s, s)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   488
"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
   489
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   490
code_pred tupled_exec .
32424
0fb428f9b5b0 New example: IMP
nipkow
parents: 32408
diff changeset
   491
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   492
subsection {* CCS *}
32408
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   493
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   494
text{* This example formalizes finite CCS processes without communication or
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   495
recursion. For simplicity, labels are natural numbers. *}
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   496
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   497
datatype proc = nil | pre nat proc | or proc proc | par proc proc
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   498
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   499
inductive step :: "proc \<Rightarrow> nat \<Rightarrow> proc \<Rightarrow> bool" where
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   500
"step (pre n p) n p" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   501
"step p1 a q \<Longrightarrow> step (or p1 p2) a q" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   502
"step p2 a q \<Longrightarrow> step (or p1 p2) a q" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   503
"step p1 a q \<Longrightarrow> step (par p1 p2) a (par q p2)" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   504
"step p2 a q \<Longrightarrow> step (par p1 p2) a (par p1 q)"
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   505
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   506
code_pred step .
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   507
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   508
inductive steps where
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   509
"steps p [] p" |
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   510
"step p a q \<Longrightarrow> steps q as r \<Longrightarrow> steps p (a#as) r"
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   511
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   512
code_pred steps .
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   513
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   514
values 5
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   515
 "{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
   516
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   517
(* FIXME
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   518
values 3 "{(a,q). step (par nil nil) a q}"
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   519
*)
a1a85b0a26f7 new interval lemma
nipkow
parents: 32355
diff changeset
   520
33105
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   521
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
   522
where
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   523
"tupled_step (pre n p, n, p)" |
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   524
"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
   525
"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
   526
"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
   527
"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
   528
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   529
code_pred tupled_step .
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   530
thm tupled_step.equation
1e4146bf841c added tupled versions of examples for the predicate compiler
bulwahn
parents: 33104
diff changeset
   531
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   532
subsection {* divmod *}
32579
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   533
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   534
inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   535
    "k < l \<Longrightarrow> divmod_rel k l 0 k"
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   536
  | "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
   537
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   538
code_pred divmod_rel ..
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   539
33111
db5af7b86a2f developing an executable the operator
bulwahn
parents: 33108
diff changeset
   540
value [code] "Predicate.the (divmod_rel_1_2 1705 42)"
32579
73ad5dbf1034 added singleton example
haftmann
parents: 32424
diff changeset
   541
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   542
subsection {* Minimum *}
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   543
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   544
definition Min
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   545
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
   546
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   547
code_pred [inductify] Min .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   548
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   549
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
   550
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   551
code_pred [inductify] lexord .
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
   552
code_pred [inductify, rpred] lexord .
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   553
thm lexord.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
   554
thm lexord.rpred_equation
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
   555
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   556
inductive less_than_nat :: "nat * nat => bool"
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   557
where
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   558
  "less_than_nat (0, x)"
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   559
| "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
   560
 
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   561
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
   562
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   563
code_pred [depth_limited] 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
   564
code_pred [rpred] less_than_nat .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   565
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   566
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
   567
where
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   568
  "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
   569
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
   570
code_pred [rpred] test_lexord .
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   571
code_pred [depth_limited] test_lexord .
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   572
thm test_lexord.depth_limited_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
   573
thm test_lexord.rpred_equation
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   574
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
   575
values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
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
   576
values [depth_limit = 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
   577
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
   578
lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   579
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   580
code_pred [inductify] lexn .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   581
thm lexn.equation
33139
9c01ee6f8ee9 added skip_proof option; playing with compilation of depth-limited predicates
bulwahn
parents: 33138
diff changeset
   582
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
   583
code_pred [rpred] lexn .
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
   584
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
   585
thm lexn.rpred_equation
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   586
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   587
code_pred [inductify, show_steps] lenlex .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   588
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
   589
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   590
code_pred [inductify, rpred] lenlex .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   591
thm lenlex.rpred_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
   592
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   593
code_pred [inductify] lists .
32670
cc0bae788b7e added a new example for the predicate compiler
bulwahn
parents: 32669
diff changeset
   594
thm lists.equation
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   595
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   596
subsection {* AVL Tree *}
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   597
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   598
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
   599
fun height :: "'a tree => nat" where
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   600
"height ET = 0"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   601
| "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
   602
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   603
consts avl :: "'a tree => bool"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   604
primrec
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   605
  "avl ET = True"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   606
  "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
   607
  h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"
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
   608
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   609
code_pred [inductify] avl .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   610
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
   611
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
   612
code_pred [rpred] avl .
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   613
thm avl.rpred_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
   614
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   615
fun set_of
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   616
where
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   617
"set_of ET = {}"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   618
| "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
   619
33123
3c7c4372f9ad cleaned up debugging messages; added options to code_pred command
bulwahn
parents: 33120
diff changeset
   620
fun is_ord :: "nat tree => bool"
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   621
where
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   622
"is_ord ET = True"
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   623
| "is_ord (MKT n l r h) =
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   624
 ((\<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
   625
33140
83951822bfd0 cleaning the signature of the predicate compiler core; renaming signature and structures to uniform long names
bulwahn
parents: 33139
diff changeset
   626
code_pred (mode: [1], [1, 2]) [inductify] set_of .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   627
thm set_of.equation
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   628
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
   629
code_pred [inductify] is_ord .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   630
thm is_ord.equation
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   631
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
   632
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   633
subsection {* Definitions about Relations *}
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   634
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   635
code_pred [inductify] converse .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   636
thm converse.equation
33145
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   637
code_pred [inductify] rel_comp .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   638
thm rel_comp.equation
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   639
code_pred [inductify] Image .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   640
thm Image.equation
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   641
(*TODO: *)
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   642
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   643
declare Id_on_def[unfolded UNION_def, code_pred_def]
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   644
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   645
code_pred [inductify] Id_on .
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   646
thm Id_on.equation
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   647
code_pred [inductify] Domain .
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   648
thm Domain.equation
33145
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   649
code_pred [inductify] Range .
33253
d9ca0c3bf680 changes to example file
bulwahn
parents: 33250
diff changeset
   650
thm Range.equation
33145
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   651
code_pred [inductify] Field .
33146
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   652
declare Sigma_def[unfolded UNION_def, code_pred_def]
bf852ef586f2 now the predicate compilere handles the predicate without introduction rules better as before
bulwahn
parents: 33145
diff changeset
   653
declare refl_on_def[unfolded UNION_def, code_pred_def]
33255
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   654
code_pred [inductify] refl_on .
75b01355e5d4 adding test case for inlining of predicate compiler
bulwahn
parents: 33253
diff changeset
   655
thm refl_on.equation
33145
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   656
code_pred [inductify] total_on .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   657
thm total_on.equation
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   658
code_pred [inductify] antisym .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   659
thm antisym.equation
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   660
code_pred [inductify] trans .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   661
thm trans.equation
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   662
code_pred [inductify] single_valued .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   663
thm single_valued.equation
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   664
code_pred [inductify] inv_image .
1a22f7ca1dfc removed dead code; added examples
bulwahn
parents: 33143
diff changeset
   665
thm inv_image.equation
32668
b2de45007537 added first prototype of the extended predicate compiler
bulwahn
parents: 32665
diff changeset
   666
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   667
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
   668
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   669
code_pred [inductify] length .
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   670
code_pred [inductify, rpred] 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
   671
thm size_listP.equation
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   672
thm size_listP.rpred_equation
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   673
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
   674
values [random] 20 "{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
   675
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   676
code_pred [inductify] concat .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   677
code_pred [inductify] hd .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   678
code_pred [inductify] tl .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   679
code_pred [inductify] last .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   680
code_pred [inductify] butlast .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   681
code_pred [inductify] take .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   682
code_pred [inductify] drop .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   683
code_pred [inductify] zip .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   684
code_pred [inductify] upt .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   685
code_pred [inductify] remdups .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   686
code_pred [inductify] remove1 .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   687
code_pred [inductify] removeAll .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   688
code_pred [inductify] distinct .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   689
code_pred [inductify] replicate .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   690
code_pred [inductify] splice .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   691
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
   692
code_pred [inductify] map .
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   693
code_pred [inductify] foldr .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   694
code_pred [inductify] foldl .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   695
code_pred [inductify] filter .
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   696
code_pred [inductify, rpred] 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
   697
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   698
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
   699
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   700
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
   701
462b1dd67a58 added context free grammar example; removed dead code; adapted to work without quick and dirty mode; fixed typo
bulwahn
parents: 32668
diff changeset
   702
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
   703
  "[] \<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
   704
| "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
   705
| "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
   706
| "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
   707
| "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
   708
| "\<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
   709
33140
83951822bfd0 cleaning the signature of the predicate compiler core; renaming signature and structures to uniform long names
bulwahn
parents: 33139
diff changeset
   710
code_pred [inductify] S\<^isub>1p .
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
   711
code_pred [inductify, rpred] 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
   712
thm S\<^isub>1p.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
   713
thm S\<^isub>1p.rpred_equation
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   714
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
   715
values [random] 5 "{x. S\<^isub>1p x}"
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   716
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   717
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
   718
  "[] \<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
   719
| "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
   720
| "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
   721
| "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
   722
| "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
   723
| "\<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
   724
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   725
code_pred [inductify, rpred] S\<^isub>2 .
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   726
thm S\<^isub>2.rpred_equation
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   727
thm A\<^isub>2.rpred_equation
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   728
thm B\<^isub>2.rpred_equation
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   729
33143
730a2e8a6ec6 modularized the compilation in the predicate compiler
bulwahn
parents: 33141
diff changeset
   730
values [random] 10 "{x. S\<^isub>2 x}"
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   731
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   732
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
   733
  "[] \<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
   734
| "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
   735
| "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
   736
| "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
   737
| "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
   738
| "\<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
   739
33124
5378e61add1a continued cleaning up; moved tuple expanding to core
bulwahn
parents: 33123
diff changeset
   740
code_pred [inductify] S\<^isub>3 .
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   741
thm S\<^isub>3.equation
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   742
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   743
values 10 "{x. S\<^isub>3 x}"
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
   744
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   745
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
   746
  "[] \<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
   747
| "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
   748
| "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
   749
| "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
   750
| "\<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
   751
| "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
   752
| "\<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
   753
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   754
code_pred (mode: [], [1]) S\<^isub>4p .
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   755
33250
5c2af18a3237 including the predicate compiler in HOL-Main; added RandomPredicate monad to Quickcheck
bulwahn
parents: 33147
diff changeset
   756
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
   757
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   758
datatype type =
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   759
    Atom nat
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   760
  | 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
   761
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   762
datatype dB =
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   763
    Var nat
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   764
  | 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
   765
  | Abs type dB
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   766
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   767
primrec
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   768
  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
   769
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   770
  "[]\<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
   771
| "(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
   772
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   773
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
   774
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   775
  "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
   776
| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
33128
1f990689349f further cleaning up
bulwahn
parents: 33126
diff changeset
   777
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   778
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
   779
  where
33128
1f990689349f further cleaning up
bulwahn
parents: 33126
diff changeset
   780
    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
   781
  | 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
   782
  | 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
   783
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   784
primrec
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   785
  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
   786
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   787
    "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
   788
  | "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
   789
  | "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
   790
32672
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   791
primrec
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   792
  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
   793
where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   794
    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
   795
      (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
   796
  | 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
   797
  | 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
   798
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   799
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
   800
  where
90f3ce5d27ae added first version of quickcheck based on the predicate compiler; added a few quickcheck examples
bulwahn
parents: 32670
diff changeset
   801
    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
   802
  | 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
   803
  | 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
   804
  | 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
   805
33326
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   806
code_pred (mode: [1, 2], [1, 2, 3]) typing .
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   807
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
   808
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   809
code_pred (mode: [1], [1, 2]) beta .
7d0288d90535 improved handling of alternative rules; added test cases for alternative rules; normalized tracing in the predicate compiler
bulwahn
parents: 33258
diff changeset
   810
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
   811
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   812
end