src/HOL/Proofs/Lambda/ListOrder.thy
author blanchet
Thu, 18 Sep 2014 16:47:40 +0200
changeset 58372 bfd497f2f4c2
parent 54295 45a5523d4a63
child 58382 2ee61d28c667
permissions -rw-r--r--
moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer) * * * made example compile again
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39157
b98909faaea8 more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
wenzelm
parents: 39126
diff changeset
     1
(*  Title:      HOL/Proofs/Lambda/ListOrder.thy
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     2
    Author:     Tobias Nipkow
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     3
    Copyright   1998 TU Muenchen
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     4
*)
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
     5
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
     6
header {* Lifting an order to lists of elements *}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
     7
58372
bfd497f2f4c2 moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
blanchet
parents: 54295
diff changeset
     8
theory ListOrder
bfd497f2f4c2 moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
blanchet
parents: 54295
diff changeset
     9
imports Old_Datatype
bfd497f2f4c2 moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
blanchet
parents: 54295
diff changeset
    10
begin
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
    11
46512
4f9f61f9b535 simplified configuration options for syntax ambiguity;
wenzelm
parents: 46506
diff changeset
    12
declare [[syntax_ambiguity_warning = false]]
39126
ee117c5b3b75 configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents: 36862
diff changeset
    13
ee117c5b3b75 configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents: 36862
diff changeset
    14
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
    15
text {*
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
    16
  Lifting an order to lists of elements, relating exactly one
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
    17
  element.
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
    18
*}
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
    19
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18257
diff changeset
    20
definition
22270
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    21
  step1 :: "('a => 'a => bool) => 'a list => 'a list => bool" where
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 18257
diff changeset
    22
  "step1 r =
22270
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    23
    (\<lambda>ys xs. \<exists>us z z' vs. xs = us @ z # vs \<and> r z' z \<and> ys =
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    24
      us @ z' # vs)"
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    25
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    26
22270
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    27
lemma step1_converse [simp]: "step1 (r^--1) = (step1 r)^--1"
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    28
  apply (unfold step1_def)
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    29
  apply (blast intro!: order_antisym)
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    30
  done
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    31
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    32
lemma in_step1_converse [iff]: "(step1 (r^--1) x y) = ((step1 r)^--1 x y)"
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    33
  apply auto
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    34
  done
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    35
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    36
lemma not_Nil_step1 [iff]: "\<not> step1 r [] xs"
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    37
  apply (unfold step1_def)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    38
  apply blast
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    39
  done
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    40
22270
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    41
lemma not_step1_Nil [iff]: "\<not> step1 r xs []"
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    42
  apply (unfold step1_def)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    43
  apply blast
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    44
  done
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    45
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    46
lemma Cons_step1_Cons [iff]:
22270
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    47
    "(step1 r (y # ys) (x # xs)) =
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    48
      (r y x \<and> xs = ys \<or> x = y \<and> step1 r ys xs)"
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    49
  apply (unfold step1_def)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    50
  apply (rule iffI)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    51
   apply (erule exE)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    52
   apply (rename_tac ts)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    53
   apply (case_tac ts)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 39157
diff changeset
    54
    apply fastforce
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    55
   apply force
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    56
  apply (erule disjE)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    57
   apply blast
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    58
  apply (blast intro: Cons_eq_appendI)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    59
  done
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    60
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    61
lemma append_step1I:
22270
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    62
  "step1 r ys xs \<and> vs = us \<or> ys = xs \<and> step1 r vs us
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    63
    ==> step1 r (ys @ vs) (xs @ us)"
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    64
  apply (unfold step1_def)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    65
  apply auto
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    66
   apply blast
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    67
  apply (blast intro: append_eq_appendI)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    68
  done
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    69
18257
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
    70
lemma Cons_step1E [elim!]:
22270
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    71
  assumes "step1 r ys (x # xs)"
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    72
    and "!!y. ys = y # xs \<Longrightarrow> r y x \<Longrightarrow> R"
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    73
    and "!!zs. ys = x # zs \<Longrightarrow> step1 r zs xs \<Longrightarrow> R"
18257
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
    74
  shows R
23464
bc2563c37b1a tuned proofs -- avoid implicit prems;
wenzelm
parents: 22270
diff changeset
    75
  using assms
18257
2124b24454dd tuned induct proofs;
wenzelm
parents: 18241
diff changeset
    76
  apply (cases ys)
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    77
   apply (simp add: step1_def)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    78
  apply blast
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    79
  done
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    80
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    81
lemma Snoc_step1_SnocD:
22270
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    82
  "step1 r (ys @ [y]) (xs @ [x])
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
    83
    ==> (step1 r ys xs \<and> y = x \<or> ys = xs \<and> r y x)"
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    84
  apply (unfold step1_def)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    85
  apply (clarify del: disjCI)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    86
  apply (rename_tac vs)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    87
  apply (rule_tac xs = vs in rev_exhaust)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    88
   apply force
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    89
  apply simp
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    90
  apply blast
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    91
  done
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    92
18241
afdba6b3e383 tuned induction proofs;
wenzelm
parents: 16417
diff changeset
    93
lemma Cons_acc_step1I [intro!]:
54295
45a5523d4a63 qualifed popular user space names
haftmann
parents: 46512
diff changeset
    94
    "Wellfounded.accp r x ==> Wellfounded.accp (step1 r) xs \<Longrightarrow> Wellfounded.accp (step1 r) (x # xs)"
45a5523d4a63 qualifed popular user space names
haftmann
parents: 46512
diff changeset
    95
  apply (induct arbitrary: xs set: Wellfounded.accp)
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    96
  apply (erule thin_rl)
23750
a1db5f819d00 - Renamed inductive2 to inductive
berghofe
parents: 23464
diff changeset
    97
  apply (erule accp_induct)
a1db5f819d00 - Renamed inductive2 to inductive
berghofe
parents: 23464
diff changeset
    98
  apply (rule accp.accI)
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
    99
  apply blast
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   100
  done
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   101
54295
45a5523d4a63 qualifed popular user space names
haftmann
parents: 46512
diff changeset
   102
lemma lists_accD: "listsp (Wellfounded.accp r) xs ==> Wellfounded.accp (step1 r) xs"
22270
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
   103
  apply (induct set: listsp)
23750
a1db5f819d00 - Renamed inductive2 to inductive
berghofe
parents: 23464
diff changeset
   104
   apply (rule accp.accI)
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   105
   apply simp
23750
a1db5f819d00 - Renamed inductive2 to inductive
berghofe
parents: 23464
diff changeset
   106
  apply (rule accp.accI)
a1db5f819d00 - Renamed inductive2 to inductive
berghofe
parents: 23464
diff changeset
   107
  apply (fast dest: accp_downward)
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   108
  done
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   109
9811
39ffdb8cab03 HOL/Lambda: converted into new-style theory and document;
wenzelm
parents: 9771
diff changeset
   110
lemma ex_step1I:
22270
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
   111
  "[| x \<in> set xs; r y x |]
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
   112
    ==> \<exists>ys. step1 r ys xs \<and> y \<in> set ys"
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   113
  apply (unfold step1_def)
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   114
  apply (drule in_set_conv_decomp [THEN iffD1])
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   115
  apply force
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   116
  done
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   117
54295
45a5523d4a63 qualifed popular user space names
haftmann
parents: 46512
diff changeset
   118
lemma lists_accI: "Wellfounded.accp (step1 r) xs ==> listsp (Wellfounded.accp r) xs"
45a5523d4a63 qualifed popular user space names
haftmann
parents: 46512
diff changeset
   119
  apply (induct set: Wellfounded.accp)
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   120
  apply clarify
23750
a1db5f819d00 - Renamed inductive2 to inductive
berghofe
parents: 23464
diff changeset
   121
  apply (rule accp.accI)
22270
4ccb7e6be929 Converted to predicate notation.
berghofe
parents: 21404
diff changeset
   122
  apply (drule_tac r=r in ex_step1I, assumption)
9771
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   123
  apply blast
54c6a2c6e569 converted Lambda scripts;
wenzelm
parents: 9102
diff changeset
   124
  done
5261
ce3c25c8a694 First steps towards termination of simply typed terms.
nipkow
parents:
diff changeset
   125
11639
wenzelm
parents: 10264
diff changeset
   126
end