src/HOL/ex/Predicate_Compile_ex.thy
author bulwahn
Thu, 11 Jun 2009 22:04:23 +0200
changeset 31575 2263d89fa930
parent 31573 0047df9eb347
child 31580 1c143f6a2226
permissions -rw-r--r--
masking proof with a quick-and-dirty step
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
d2cead76fca2 split Predicate_Compile examples into separate theory
haftmann
parents: 31123
diff changeset
     2
imports Complex_Main Predicate_Compile
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
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
     5
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
     6
    "even 0"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
     7
  | "even n \<Longrightarrow> odd (Suc n)"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
     8
  | "odd n \<Longrightarrow> even (Suc n)"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
     9
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
    10
code_pred even .
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
    11
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
    12
thm odd.equation
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
    13
thm even.equation
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    14
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    15
values "{x. even 2}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    16
values "{x. odd 2}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    17
values 10 "{n. even n}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    18
values 10 "{n. odd n}"
31195
12741f23527d added example on ML level
haftmann
parents: 31129
diff changeset
    19
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    20
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    21
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    22
    append_Nil: "append [] xs xs"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    23
  | append_Cons: "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    24
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
    25
inductive rev
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
    26
where
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
    27
"rev [] []"
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
    28
| "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys"
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
    29
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
    30
code_pred rev .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
    31
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
    32
thm append.equation
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    33
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    34
values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    35
values "{zs. append [0, Suc 0, 2] [17, 8] zs}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    36
values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0,5]}"
31195
12741f23527d added example on ML level
haftmann
parents: 31129
diff changeset
    37
30972
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    38
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    39
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
    40
  for f where
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    41
    "partition f [] [] []"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    42
  | "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
    43
  | "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"
5b65835ccc92 some experiements towards user interface for predicate compiler
haftmann
parents: 30953
diff changeset
    44
31551
995d6b90e9d6 making isatest happy; but misunderstanding remains
bulwahn
parents: 31550
diff changeset
    45
(* FIXME: correct handling of parameters *)
995d6b90e9d6 making isatest happy; but misunderstanding remains
bulwahn
parents: 31550
diff changeset
    46
(*
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
    47
ML {* reset Predicate_Compile.do_proofs *}
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
    48
code_pred partition .
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
    49
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
    50
thm partition.equation
31551
995d6b90e9d6 making isatest happy; but misunderstanding remains
bulwahn
parents: 31550
diff changeset
    51
ML {* set Predicate_Compile.do_proofs *}
995d6b90e9d6 making isatest happy; but misunderstanding remains
bulwahn
parents: 31550
diff changeset
    52
*)
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
    53
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
    54
(* TODO: requires to handle abstractions in parameter positions correctly *)
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    55
(*FIXME values 10 "{(ys, zs). partition (\<lambda>n. n mod 2 = 0)
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
    56
  [0, Suc 0, 2, 3, 4, 5, 6, 7] ys zs}" *)
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
    57
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    58
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
    59
lemma [code_pred_intros]:
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
    60
"r a b ==> tranclp r a b"
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
    61
"r a b ==> tranclp r b c ==> tranclp r a c" 
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
    62
by auto
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    63
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    64
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    65
lemma converse_tranclpE:
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    66
  assumes "tranclp r x z "
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    67
  assumes "r x z ==> P"
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    68
  assumes "\<And> y. [| r x y; tranclp r y z |] ==> P"
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    69
  shows P
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    70
proof -
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    71
  from tranclpD[OF assms(1)]
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    72
  obtain y where "r x y" and "rtranclp r y z" by iprover
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    73
  with assms(2-3) rtranclpD[OF this(2)] this(1)
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    74
  show P by iprover
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    75
qed  
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    76
31575
2263d89fa930 masking proof with a quick-and-dirty step
bulwahn
parents: 31573
diff changeset
    77
(* Setup requires quick and dirty proof *)
2263d89fa930 masking proof with a quick-and-dirty step
bulwahn
parents: 31573
diff changeset
    78
(*
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    79
code_pred tranclp
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    80
proof -
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    81
  assume tranclp: "tranclp r a1 a2"
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    82
     and step: "\<And> a c b. a1 = a ==> a2 = c ==> r a b ==> tranclp r b c ==> thesis"
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    83
     and base: "\<And> a b. a1 = a ==> a2 = b ==> r a b ==> thesis"
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    84
  show thesis
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    85
  proof (cases rule: converse_tranclpE[OF tranclp])
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    86
    case 1
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    87
    from 1 base show thesis by auto
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    88
  next
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    89
    case 2
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    90
    from 2 step show thesis by auto
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    91
  qed
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    92
qed
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
    93
31123
e3b4e52c01c2 examples using code_pred
haftmann
parents: 31111
diff changeset
    94
thm tranclp.equation
31575
2263d89fa930 masking proof with a quick-and-dirty step
bulwahn
parents: 31573
diff changeset
    95
*)
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
    96
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    97
inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    98
    "succ 0 1"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
    99
  | "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   100
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
   101
code_pred succ .
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
   102
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   103
thm succ.equation
31573
0047df9eb347 improved infrastructure of predicate compiler for adding manual introduction rules
bulwahn
parents: 31551
diff changeset
   104
(* FIXME: why does this not terminate? *)
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
   105
(*
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   106
values 20 "{n. tranclp succ 10 n}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   107
values "{n. tranclp succ n 10}"
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   108
values 20 "{(n, m). tranclp succ n m}"
31514
fed8a95f54db refactoring the predicate compiler
bulwahn
parents: 31217
diff changeset
   109
*)
31217
c025f32afd4e experimental values command
haftmann
parents: 31195
diff changeset
   110
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
   111
end