src/HOL/ex/Induction_Schema.thy
author hoelzl
Wed, 01 Dec 2010 20:12:53 +0100
changeset 40870 94427db32392
parent 33471 5aef13872723
child 58889 5b7a9633cfa8
permissions -rw-r--r--
Tuned setup for borel_measurable with min, max and psuminf.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33471
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     1
(*  Title:      HOL/ex/Induction_Schema.thy
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     2
    Author:     Alexander Krauss, TU Muenchen
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     3
*)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     4
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     5
header {* Examples of automatically derived induction rules *}
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     6
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     7
theory Induction_Schema
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     8
imports Main
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
     9
begin
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    10
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    11
subsection {* Some simple induction principles on nat *}
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    12
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    13
lemma nat_standard_induct: (* cf. Nat.thy *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    14
  "\<lbrakk>P 0; \<And>n. P n \<Longrightarrow> P (Suc n)\<rbrakk> \<Longrightarrow> P x"
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    15
by induction_schema (pat_completeness, lexicographic_order)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    16
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    17
lemma nat_induct2:
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    18
  "\<lbrakk> P 0; P (Suc 0); \<And>k. P k ==> P (Suc k) ==> P (Suc (Suc k)) \<rbrakk>
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    19
  \<Longrightarrow> P n"
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    20
by induction_schema (pat_completeness, lexicographic_order)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    21
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    22
lemma minus_one_induct:
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    23
  "\<lbrakk>\<And>n::nat. (n \<noteq> 0 \<Longrightarrow> P (n - 1)) \<Longrightarrow> P n\<rbrakk> \<Longrightarrow> P x"
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    24
by induction_schema (pat_completeness, lexicographic_order)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    25
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    26
theorem diff_induct: (* cf. Nat.thy *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    27
  "(!!x. P x 0) ==> (!!y. P 0 (Suc y)) ==>
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    28
    (!!x y. P x y ==> P (Suc x) (Suc y)) ==> P m n"
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    29
by induction_schema (pat_completeness, lexicographic_order)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    30
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    31
lemma list_induct2': (* cf. List.thy *)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    32
  "\<lbrakk> P [] [];
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    33
  \<And>x xs. P (x#xs) [];
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    34
  \<And>y ys. P [] (y#ys);
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    35
   \<And>x xs y ys. P xs ys  \<Longrightarrow> P (x#xs) (y#ys) \<rbrakk>
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    36
 \<Longrightarrow> P xs ys"
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    37
by induction_schema (pat_completeness, lexicographic_order)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    38
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    39
theorem even_odd_induct:
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    40
  assumes "R 0"
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    41
  assumes "Q 0"
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    42
  assumes "\<And>n. Q n \<Longrightarrow> R (Suc n)"
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    43
  assumes "\<And>n. R n \<Longrightarrow> Q (Suc n)"
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    44
  shows "R n" "Q n"
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    45
  using assms
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    46
by induction_schema (pat_completeness+, lexicographic_order)
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    47
5aef13872723 renamed method induct_scheme to induction_schema
krauss
parents:
diff changeset
    48
end