src/HOL/UNITY/ListOrder.thy
author Fabian Huch <huch@in.tum.de>
Wed, 18 Oct 2023 20:12:07 +0200
changeset 78843 fc3ba0a1c82f
parent 67613 ce654b0e6d69
child 80914 d97fdabd9e2b
permissions -rw-r--r--
read relative cpu from build log;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32235
diff changeset
     1
(*  Title:      HOL/UNITY/ListOrder.thy
6708
62beb3336b02 lists are partially ordered by the prefix relation
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
62beb3336b02 lists are partially ordered by the prefix relation
paulson
parents:
diff changeset
     3
    Copyright   1998  University of Cambridge
62beb3336b02 lists are partially ordered by the prefix relation
paulson
parents:
diff changeset
     4
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
     5
Lists are partially ordered by Charpentier's Generalized Prefix Relation
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
     6
   (xs,ys) : genPrefix(r)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
     7
     if ys = xs' @ zs where length xs = length xs'
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
     8
     and corresponding elements of xs, xs' are pairwise related by r
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
     9
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    10
Also overloads <= and < for lists!
6708
62beb3336b02 lists are partially ordered by the prefix relation
paulson
parents:
diff changeset
    11
*)
62beb3336b02 lists are partially ordered by the prefix relation
paulson
parents:
diff changeset
    12
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    13
section \<open>The Prefix Ordering on Lists\<close>
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    14
27682
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    15
theory ListOrder
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    16
imports Main
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    17
begin
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    18
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    19
inductive_set
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    20
  genPrefix :: "('a * 'a)set => ('a list * 'a list)set"
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    21
  for r :: "('a * 'a)set"
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    22
 where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    23
   Nil:     "([],[]) \<in> genPrefix(r)"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    24
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    25
 | prepend: "[| (xs,ys) \<in> genPrefix(r);  (x,y) \<in> r |] ==>
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    26
             (x#xs, y#ys) \<in> genPrefix(r)"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    27
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    28
 | append:  "(xs,ys) \<in> genPrefix(r) ==> (xs, ys@zs) \<in> genPrefix(r)"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    29
27682
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    30
instantiation list :: (type) ord 
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    31
begin
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    32
27682
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    33
definition
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    34
  prefix_def:        "xs <= zs \<longleftrightarrow>  (xs, zs) \<in> genPrefix Id"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    35
27682
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    36
definition
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    37
  strict_prefix_def: "xs < zs  \<longleftrightarrow>  xs \<le> zs \<and> \<not> zs \<le> (xs :: 'a list)"
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    38
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    39
instance ..  
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    40
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    41
(*Constants for the <= and >= relations, used below in translations*)
27682
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    42
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    43
end
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
    44
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    45
definition Le :: "(nat*nat) set" where
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    46
    "Le == {(x,y). x <= y}"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    47
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 32960
diff changeset
    48
definition  Ge :: "(nat*nat) set" where
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    49
    "Ge == {(x,y). y <= x}"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    50
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    51
abbreviation
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    52
  pfixLe :: "[nat list, nat list] => bool"  (infixl "pfixLe" 50)  where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    53
  "xs pfixLe ys == (xs,ys) \<in> genPrefix Le"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    54
23767
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    55
abbreviation
7272a839ccd9 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    56
  pfixGe :: "[nat list, nat list] => bool"  (infixl "pfixGe" 50)  where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    57
  "xs pfixGe ys == (xs,ys) \<in> genPrefix Ge"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    58
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    59
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    60
subsection\<open>preliminary lemmas\<close>
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    61
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    62
lemma Nil_genPrefix [iff]: "([], xs) \<in> genPrefix r"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    63
by (cut_tac genPrefix.Nil [THEN genPrefix.append], auto)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    64
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    65
lemma genPrefix_length_le: "(xs,ys) \<in> genPrefix r \<Longrightarrow> length xs <= length ys"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    66
by (erule genPrefix.induct, auto)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    67
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    68
lemma cdlemma:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    69
     "[| (xs', ys') \<in> genPrefix r |]  
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    70
      ==> (\<forall>x xs. xs' = x#xs \<longrightarrow> (\<exists>y ys. ys' = y#ys & (x,y) \<in> r & (xs, ys) \<in> genPrefix r))"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    71
apply (erule genPrefix.induct, blast, blast)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    72
apply (force intro: genPrefix.append)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    73
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    74
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    75
(*As usual converting it to an elimination rule is tiresome*)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    76
lemma cons_genPrefixE [elim!]: 
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    77
     "[| (x#xs, zs) \<in> genPrefix r;   
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    78
         !!y ys. [| zs = y#ys;  (x,y) \<in> r;  (xs, ys) \<in> genPrefix r |] ==> P  
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    79
      |] ==> P"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    80
by (drule cdlemma, simp, blast)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    81
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    82
lemma Cons_genPrefix_Cons [iff]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    83
     "((x#xs,y#ys) \<in> genPrefix r) = ((x,y) \<in> r \<and> (xs,ys) \<in> genPrefix r)"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    84
by (blast intro: genPrefix.prepend)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    85
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    86
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    87
subsection\<open>genPrefix is a partial order\<close>
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    88
30198
922f944f03b2 name changes
nipkow
parents: 27682
diff changeset
    89
lemma refl_genPrefix: "refl r ==> refl (genPrefix r)"
922f944f03b2 name changes
nipkow
parents: 27682
diff changeset
    90
apply (unfold refl_on_def, auto)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    91
apply (induct_tac "x")
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    92
prefer 2 apply (blast intro: genPrefix.prepend)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    93
apply (blast intro: genPrefix.Nil)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    94
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    95
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
    96
lemma genPrefix_refl [simp]: "refl r \<Longrightarrow> (l,l) \<in> genPrefix r"
30198
922f944f03b2 name changes
nipkow
parents: 27682
diff changeset
    97
by (erule refl_onD [OF refl_genPrefix UNIV_I])
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    98
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
    99
lemma genPrefix_mono: "r<=s ==> genPrefix r <= genPrefix s"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   100
apply clarify
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   101
apply (erule genPrefix.induct)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   102
apply (auto intro: genPrefix.append)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   103
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   104
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   105
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   106
(** Transitivity **)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   107
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   108
(*A lemma for proving genPrefix_trans_O*)
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   109
lemma append_genPrefix:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   110
     "(xs @ ys, zs) \<in> genPrefix r \<Longrightarrow> (xs, zs) \<in> genPrefix r"
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   111
  by (induct xs arbitrary: zs) auto
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   112
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   113
(*Lemma proving transitivity and more*)
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   114
lemma genPrefix_trans_O:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   115
  assumes "(x, y) \<in> genPrefix r"
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   116
  shows "\<And>z. (y, z) \<in> genPrefix s \<Longrightarrow> (x, z) \<in> genPrefix (r O s)"
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   117
  apply (atomize (full))
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   118
  using assms
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   119
  apply induct
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   120
    apply blast
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   121
   apply (blast intro: genPrefix.prepend)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   122
  apply (blast dest: append_genPrefix)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   123
  done
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   124
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   125
lemma genPrefix_trans:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   126
  "(x, y) \<in> genPrefix r \<Longrightarrow> (y, z) \<in> genPrefix r \<Longrightarrow> trans r
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   127
    \<Longrightarrow> (x, z) \<in> genPrefix r"
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   128
  apply (rule trans_O_subset [THEN genPrefix_mono, THEN subsetD])
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   129
   apply assumption
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   130
  apply (blast intro: genPrefix_trans_O)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   131
  done
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   132
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   133
lemma prefix_genPrefix_trans:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   134
  "[| x<=y;  (y,z) \<in> genPrefix r |] ==> (x, z) \<in> genPrefix r"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   135
apply (unfold prefix_def)
32235
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 30198
diff changeset
   136
apply (drule genPrefix_trans_O, assumption)
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 30198
diff changeset
   137
apply simp
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   138
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   139
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   140
lemma genPrefix_prefix_trans:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   141
  "[| (x,y) \<in> genPrefix r;  y<=z |] ==> (x,z) \<in> genPrefix r"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   142
apply (unfold prefix_def)
32235
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 30198
diff changeset
   143
apply (drule genPrefix_trans_O, assumption)
8f9b8d14fc9f "more standard" argument order of relation composition (op O)
krauss
parents: 30198
diff changeset
   144
apply simp
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   145
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   146
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   147
lemma trans_genPrefix: "trans r ==> trans (genPrefix r)"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   148
by (blast intro: transI genPrefix_trans)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   149
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   150
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   151
(** Antisymmetry **)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   152
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   153
lemma genPrefix_antisym:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   154
  assumes 1: "(xs, ys) \<in> genPrefix r"
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   155
    and 2: "antisym r"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   156
    and 3: "(ys, xs) \<in> genPrefix r"
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   157
  shows "xs = ys"
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   158
  using 1 3
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   159
proof induct
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   160
  case Nil
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   161
  then show ?case by blast
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   162
next
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   163
  case prepend
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   164
  then show ?case using 2 by (simp add: antisym_def)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   165
next
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   166
  case (append xs ys zs)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   167
  then show ?case
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   168
    apply -
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   169
    apply (subgoal_tac "length zs = 0", force)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   170
    apply (drule genPrefix_length_le)+
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   171
    apply (simp del: length_0_conv)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   172
    done
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   173
qed
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   174
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   175
lemma antisym_genPrefix: "antisym r ==> antisym (genPrefix r)"
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   176
  by (blast intro: antisymI genPrefix_antisym)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   177
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   178
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   179
subsection\<open>recursion equations\<close>
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   180
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   181
lemma genPrefix_Nil [simp]: "((xs, []) \<in> genPrefix r) = (xs = [])"
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 46577
diff changeset
   182
  by (induct xs) auto
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   183
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   184
lemma same_genPrefix_genPrefix [simp]: 
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   185
    "refl r \<Longrightarrow> ((xs@ys, xs@zs) \<in> genPrefix r) = ((ys,zs) \<in> genPrefix r)"
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 46577
diff changeset
   186
  by (induct xs) (simp_all add: refl_on_def)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   187
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   188
lemma genPrefix_Cons:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   189
     "((xs, y#ys) \<in> genPrefix r) =  
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   190
      (xs=[] | (\<exists>z zs. xs=z#zs & (z,y) \<in> r & (zs,ys) \<in> genPrefix r))"
46911
6d2a2f0e904e tuned proofs;
wenzelm
parents: 46577
diff changeset
   191
  by (cases xs) auto
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   192
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   193
lemma genPrefix_take_append:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   194
     "[| refl r;  (xs,ys) \<in> genPrefix r |]  
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   195
      ==>  (xs@zs, take (length xs) ys @ zs) \<in> genPrefix r"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   196
apply (erule genPrefix.induct)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   197
apply (frule_tac [3] genPrefix_length_le)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   198
apply (simp_all (no_asm_simp) add: diff_is_0_eq [THEN iffD2])
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   199
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   200
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   201
lemma genPrefix_append_both:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   202
     "[| refl r;  (xs,ys) \<in> genPrefix r;  length xs = length ys |]  
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   203
      ==>  (xs@zs, ys @ zs) \<in> genPrefix r"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   204
apply (drule genPrefix_take_append, assumption)
46577
e5438c5797ae tuned proofs;
wenzelm
parents: 45477
diff changeset
   205
apply simp
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   206
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   207
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   208
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   209
(*NOT suitable for rewriting since [y] has the form y#ys*)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   210
lemma append_cons_eq: "xs @ y # ys = (xs @ [y]) @ ys"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   211
by auto
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   212
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   213
lemma aolemma:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   214
     "[| (xs,ys) \<in> genPrefix r;  refl r |]  
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   215
      ==> length xs < length ys \<longrightarrow> (xs @ [ys ! length xs], ys) \<in> genPrefix r"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   216
apply (erule genPrefix.induct)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   217
  apply blast
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   218
 apply simp
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   219
txt\<open>Append case is hardest\<close>
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   220
apply simp
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   221
apply (frule genPrefix_length_le [THEN le_imp_less_or_eq])
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   222
apply (erule disjE)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   223
apply (simp_all (no_asm_simp) add: neq_Nil_conv nth_append)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   224
apply (blast intro: genPrefix.append, auto)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13798
diff changeset
   225
apply (subst append_cons_eq, fast intro: genPrefix_append_both genPrefix.append)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   226
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   227
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   228
lemma append_one_genPrefix:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   229
     "[| (xs,ys) \<in> genPrefix r;  length xs < length ys;  refl r |]  
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   230
      ==> (xs @ [ys ! length xs], ys) \<in> genPrefix r"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   231
by (blast intro: aolemma [THEN mp])
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   232
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   233
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   234
(** Proving the equivalence with Charpentier's definition **)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   235
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   236
lemma genPrefix_imp_nth:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   237
    "i < length xs \<Longrightarrow> (xs, ys) \<in> genPrefix r \<Longrightarrow> (xs ! i, ys ! i) \<in> r"
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   238
  apply (induct xs arbitrary: i ys)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   239
   apply auto
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   240
  apply (case_tac i)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   241
   apply auto
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   242
  done
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   243
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   244
lemma nth_imp_genPrefix:
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   245
  "length xs <= length ys \<Longrightarrow>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   246
     (\<forall>i. i < length xs \<longrightarrow> (xs ! i, ys ! i) \<in> r) \<Longrightarrow>
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   247
     (xs, ys) \<in> genPrefix r"
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   248
  apply (induct xs arbitrary: ys)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   249
   apply (simp_all add: less_Suc_eq_0_disj all_conj_distrib)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   250
  apply (case_tac ys)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   251
   apply (force+)
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   252
  done
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   253
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   254
lemma genPrefix_iff_nth:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   255
     "((xs,ys) \<in> genPrefix r) =  
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   256
      (length xs <= length ys & (\<forall>i. i < length xs \<longrightarrow> (xs!i, ys!i) \<in> r))"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   257
apply (blast intro: genPrefix_length_le genPrefix_imp_nth nth_imp_genPrefix)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   258
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   259
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   260
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   261
subsection\<open>The type of lists is partially ordered\<close>
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   262
30198
922f944f03b2 name changes
nipkow
parents: 27682
diff changeset
   263
declare refl_Id [iff] 
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   264
        antisym_Id [iff] 
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   265
        trans_Id [iff]
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   266
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   267
lemma prefix_refl [iff]: "xs <= (xs::'a list)"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   268
by (simp add: prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   269
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   270
lemma prefix_trans: "!!xs::'a list. [| xs <= ys; ys <= zs |] ==> xs <= zs"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   271
apply (unfold prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   272
apply (blast intro: genPrefix_trans)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   273
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   274
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   275
lemma prefix_antisym: "!!xs::'a list. [| xs <= ys; ys <= xs |] ==> xs = ys"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   276
apply (unfold prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   277
apply (blast intro: genPrefix_antisym)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   278
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   279
27682
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
   280
lemma prefix_less_le_not_le: "!!xs::'a list. (xs < zs) = (xs <= zs & \<not> zs \<le> xs)"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   281
by (unfold strict_prefix_def, auto)
6708
62beb3336b02 lists are partially ordered by the prefix relation
paulson
parents:
diff changeset
   282
12338
de0f4a63baa5 renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents: 6810
diff changeset
   283
instance list :: (type) order
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   284
  by (intro_classes,
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   285
      (assumption | rule prefix_refl prefix_trans prefix_antisym
27682
25aceefd4786 added class preorder
haftmann
parents: 23767
diff changeset
   286
                     prefix_less_le_not_le)+)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   287
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   288
(*Monotonicity of "set" operator WRT prefix*)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   289
lemma set_mono: "xs <= ys ==> set xs <= set ys"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   290
apply (unfold prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   291
apply (erule genPrefix.induct, auto)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   292
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   293
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   294
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   295
(** recursion equations **)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   296
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   297
lemma Nil_prefix [iff]: "[] <= xs"
46577
e5438c5797ae tuned proofs;
wenzelm
parents: 45477
diff changeset
   298
by (simp add: prefix_def)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   299
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   300
lemma prefix_Nil [simp]: "(xs <= []) = (xs = [])"
46577
e5438c5797ae tuned proofs;
wenzelm
parents: 45477
diff changeset
   301
by (simp add: prefix_def)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   302
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   303
lemma Cons_prefix_Cons [simp]: "(x#xs <= y#ys) = (x=y & xs<=ys)"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   304
by (simp add: prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   305
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   306
lemma same_prefix_prefix [simp]: "(xs@ys <= xs@zs) = (ys <= zs)"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   307
by (simp add: prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   308
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   309
lemma append_prefix [iff]: "(xs@ys <= xs) = (ys <= [])"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   310
by (insert same_prefix_prefix [of xs ys "[]"], simp)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   311
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   312
lemma prefix_appendI [simp]: "xs <= ys ==> xs <= ys@zs"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   313
apply (unfold prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   314
apply (erule genPrefix.append)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   315
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   316
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   317
lemma prefix_Cons: 
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   318
   "(xs <= y#ys) = (xs=[] | (\<exists>zs. xs=y#zs \<and> zs <= ys))"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   319
by (simp add: prefix_def genPrefix_Cons)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   320
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   321
lemma append_one_prefix: 
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   322
  "[| xs <= ys; length xs < length ys |] ==> xs @ [ys ! length xs] <= ys"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   323
apply (unfold prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   324
apply (simp add: append_one_genPrefix)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   325
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   326
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   327
lemma prefix_length_le: "xs <= ys ==> length xs <= length ys"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   328
apply (unfold prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   329
apply (erule genPrefix_length_le)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   330
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   331
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   332
lemma splemma: "xs<=ys ==> xs~=ys --> length xs < length ys"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   333
apply (unfold prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   334
apply (erule genPrefix.induct, auto)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   335
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   336
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   337
lemma strict_prefix_length_less: "xs < ys ==> length xs < length ys"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   338
apply (unfold strict_prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   339
apply (blast intro: splemma [THEN mp])
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   340
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   341
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   342
lemma mono_length: "mono length"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   343
by (blast intro: monoI prefix_length_le)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   344
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   345
(*Equivalence to the definition used in Lex/Prefix.thy*)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   346
lemma prefix_iff: "(xs <= zs) = (\<exists>ys. zs = xs@ys)"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   347
apply (unfold prefix_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   348
apply (auto simp add: genPrefix_iff_nth nth_append)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   349
apply (rule_tac x = "drop (length xs) zs" in exI)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   350
apply (rule nth_equalityI)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   351
apply (simp_all (no_asm_simp) add: nth_append)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   352
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   353
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   354
lemma prefix_snoc [simp]: "(xs <= ys@[y]) = (xs = ys@[y] | xs <= ys)"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   355
apply (simp add: prefix_iff)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   356
apply (rule iffI)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   357
 apply (erule exE)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   358
 apply (rename_tac "zs")
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   359
 apply (rule_tac xs = zs in rev_exhaust)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   360
  apply simp
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   361
 apply clarify
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   362
 apply (simp del: append_assoc add: append_assoc [symmetric], force)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   363
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   364
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   365
lemma prefix_append_iff:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 63146
diff changeset
   366
     "(xs <= ys@zs) = (xs <= ys | (\<exists>us. xs = ys@us & us <= zs))"
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   367
apply (rule_tac xs = zs in rev_induct)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   368
 apply force
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   369
apply (simp del: append_assoc add: append_assoc [symmetric], force)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   370
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   371
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   372
(*Although the prefix ordering is not linear, the prefixes of a list
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   373
  are linearly ordered.*)
45477
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   374
lemma common_prefix_linear:
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   375
  fixes xs ys zs :: "'a list"
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   376
  shows "xs <= zs \<Longrightarrow> ys <= zs \<Longrightarrow> xs <= ys | ys <= xs"
11d9c2768729 tuned proofs;
wenzelm
parents: 35416
diff changeset
   377
  by (induct zs rule: rev_induct) auto
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   378
63146
f1ecba0272f9 isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   379
subsection\<open>pfixLe, pfixGe: properties inherited from the translations\<close>
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   380
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   381
(** pfixLe **)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   382
30198
922f944f03b2 name changes
nipkow
parents: 27682
diff changeset
   383
lemma refl_Le [iff]: "refl Le"
922f944f03b2 name changes
nipkow
parents: 27682
diff changeset
   384
by (unfold refl_on_def Le_def, auto)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   385
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   386
lemma antisym_Le [iff]: "antisym Le"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   387
by (unfold antisym_def Le_def, auto)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   388
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   389
lemma trans_Le [iff]: "trans Le"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   390
by (unfold trans_def Le_def, auto)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   391
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   392
lemma pfixLe_refl [iff]: "x pfixLe x"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   393
by simp
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   394
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   395
lemma pfixLe_trans: "[| x pfixLe y; y pfixLe z |] ==> x pfixLe z"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   396
by (blast intro: genPrefix_trans)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   397
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   398
lemma pfixLe_antisym: "[| x pfixLe y; y pfixLe x |] ==> x = y"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   399
by (blast intro: genPrefix_antisym)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   400
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   401
lemma prefix_imp_pfixLe: "xs<=ys ==> xs pfixLe ys"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   402
apply (unfold prefix_def Le_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   403
apply (blast intro: genPrefix_mono [THEN [2] rev_subsetD])
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   404
done
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   405
30198
922f944f03b2 name changes
nipkow
parents: 27682
diff changeset
   406
lemma refl_Ge [iff]: "refl Ge"
922f944f03b2 name changes
nipkow
parents: 27682
diff changeset
   407
by (unfold refl_on_def Ge_def, auto)
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   408
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   409
lemma antisym_Ge [iff]: "antisym Ge"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   410
by (unfold antisym_def Ge_def, auto)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   411
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   412
lemma trans_Ge [iff]: "trans Ge"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   413
by (unfold trans_def Ge_def, auto)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   414
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   415
lemma pfixGe_refl [iff]: "x pfixGe x"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   416
by simp
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   417
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   418
lemma pfixGe_trans: "[| x pfixGe y; y pfixGe z |] ==> x pfixGe z"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   419
by (blast intro: genPrefix_trans)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   420
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   421
lemma pfixGe_antisym: "[| x pfixGe y; y pfixGe x |] ==> x = y"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   422
by (blast intro: genPrefix_antisym)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   423
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   424
lemma prefix_imp_pfixGe: "xs<=ys ==> xs pfixGe ys"
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   425
apply (unfold prefix_def Ge_def)
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   426
apply (blast intro: genPrefix_mono [THEN [2] rev_subsetD])
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 12338
diff changeset
   427
done
6708
62beb3336b02 lists are partially ordered by the prefix relation
paulson
parents:
diff changeset
   428
62beb3336b02 lists are partially ordered by the prefix relation
paulson
parents:
diff changeset
   429
end