src/HOL/HOLCF/IOA/meta_theory/Sequence.thy
author wenzelm
Wed, 30 Dec 2015 21:57:52 +0100
changeset 62002 f1599e98c4d0
parent 62001 1f2788fb0b8b
child 62005 68db98c2cd97
permissions -rw-r--r--
isabelle update_cartouches -c -t;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41476
diff changeset
     1
(*  Title:      HOL/HOLCF/IOA/meta_theory/Sequence.thy
40945
b8703f63bfb2 recoded latin1 as utf8;
wenzelm
parents: 40774
diff changeset
     2
    Author:     Olaf Müller
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     3
12218
wenzelm
parents: 12114
diff changeset
     4
Sequences over flat domains with lifted elements.
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     5
*)
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     6
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     7
theory Sequence
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
imports Seq
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
begin
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    10
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35907
diff changeset
    11
default_sort type
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    12
41476
0fa9629aa399 types -> type_synonym
huffman
parents: 41429
diff changeset
    13
type_synonym 'a Seq = "'a lift seq"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    14
3952
dca1bce88ec8 replaced ops by consts;
wenzelm
parents: 3275
diff changeset
    15
consts
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    16
7229
6773ba0c36d5 renamed Cons to Consq in order to avoid clash with List.Cons;
wenzelm
parents: 4559
diff changeset
    17
  Consq            ::"'a            => 'a Seq -> 'a Seq"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    18
  Filter           ::"('a => bool)  => 'a Seq -> 'a Seq"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    19
  Map              ::"('a => 'b)    => 'a Seq -> 'b Seq"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    20
  Forall           ::"('a => bool)  => 'a Seq => bool"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    21
  Last             ::"'a Seq        -> 'a lift"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    22
  Dropwhile        ::"('a => bool)  => 'a Seq -> 'a Seq"
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    23
  Takewhile        ::"('a => bool)  => 'a Seq -> 'a Seq"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    24
  Zip              ::"'a Seq        -> 'b Seq -> ('a * 'b) Seq"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    25
  Flat             ::"('a Seq) seq   -> 'a Seq"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    26
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    27
  Filter2          ::"('a => bool)  => 'a Seq -> 'a Seq"
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    28
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23778
diff changeset
    29
abbreviation
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
    30
  Consq_syn  ("(_/\<leadsto>_)"  [66,65] 65) where
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
    31
  "a\<leadsto>s == Consq a$s"
4283
92707e24b62b managed merge details;
mueller
parents: 3952
diff changeset
    32
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    33
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23778
diff changeset
    34
(* list Enumeration *)
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23778
diff changeset
    35
syntax
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23778
diff changeset
    36
  "_totlist"      :: "args => 'a Seq"              ("[(_)!]")
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 23778
diff changeset
    37
  "_partlist"     :: "args => 'a Seq"              ("[(_)?]")
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    38
translations
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
    39
  "[x, xs!]"     == "x\<leadsto>[xs!]"
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
    40
  "[x!]"         == "x\<leadsto>nil"
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
    41
  "[x, xs?]"     == "x\<leadsto>[xs?]"
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
    42
  "[x?]"         == "x\<leadsto>CONST bottom"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    43
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    44
defs
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    45
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    46
Consq_def:     "Consq a == LAM s. Def a ## s"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    47
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    48
Filter_def:    "Filter P == sfilter$(flift2 P)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    49
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    50
Map_def:       "Map f  == smap$(flift2 f)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    51
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    52
Forall_def:    "Forall P == sforall (flift2 P)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    53
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    54
Last_def:      "Last == slast"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    55
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    56
Dropwhile_def: "Dropwhile P == sdropwhile$(flift2 P)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    57
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    58
Takewhile_def: "Takewhile P == stakewhile$(flift2 P)"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    59
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    60
Flat_def:      "Flat == sflat"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    61
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    62
Zip_def:
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    63
  "Zip == (fix$(LAM h t1 t2. case t1 of
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    64
               nil   => nil
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    65
             | x##xs => (case t2 of
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    66
                          nil => UU
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    67
                        | y##ys => (case x of
12028
52aa183c15bb replaced Undef by UU;
wenzelm
parents: 10835
diff changeset
    68
                                      UU  => UU
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    69
                                    | Def a => (case y of
12028
52aa183c15bb replaced Undef by UU;
wenzelm
parents: 10835
diff changeset
    70
                                                  UU => UU
10835
nipkow
parents: 7229
diff changeset
    71
                                                | Def b => Def (a,b)##(h$xs$ys))))))"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    72
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    73
Filter2_def:    "Filter2 P == (fix$(LAM h t. case t of
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    74
            nil => nil
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    75
          | x##xs => (case x of UU => UU | Def y => (if P y
10835
nipkow
parents: 7229
diff changeset
    76
                     then x##(h$xs)
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    77
                     else     h$xs))))"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    78
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    79
declare andalso_and [simp]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    80
declare andalso_or [simp]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    81
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    82
subsection "recursive equations of operators"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    83
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    84
subsubsection "Map"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    85
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    86
lemma Map_UU: "Map f$UU =UU"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    87
by (simp add: Map_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    88
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    89
lemma Map_nil: "Map f$nil =nil"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    90
by (simp add: Map_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    91
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
    92
lemma Map_cons: "Map f$(x\<leadsto>xs)=(f x) \<leadsto> Map f$xs"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    93
by (simp add: Map_def Consq_def flift2_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    94
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    95
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
    96
subsubsection \<open>Filter\<close>
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    97
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    98
lemma Filter_UU: "Filter P$UU =UU"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
    99
by (simp add: Filter_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   100
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   101
lemma Filter_nil: "Filter P$nil =nil"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   102
by (simp add: Filter_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   103
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   104
lemma Filter_cons:
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   105
  "Filter P$(x\<leadsto>xs)= (if P x then x\<leadsto>(Filter P$xs) else Filter P$xs)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   106
by (simp add: Filter_def Consq_def flift2_def If_and_if)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   107
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   108
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
   109
subsubsection \<open>Forall\<close>
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   110
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   111
lemma Forall_UU: "Forall P UU"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   112
by (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   113
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   114
lemma Forall_nil: "Forall P nil"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   115
by (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   116
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   117
lemma Forall_cons: "Forall P (x\<leadsto>xs)= (P x & Forall P xs)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   118
by (simp add: Forall_def sforall_def Consq_def flift2_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   119
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   120
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
   121
subsubsection \<open>Conc\<close>
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   122
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   123
lemma Conc_cons: "(x\<leadsto>xs) @@ y = x\<leadsto>(xs @@y)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   124
by (simp add: Consq_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   125
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   126
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
   127
subsubsection \<open>Takewhile\<close>
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   128
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   129
lemma Takewhile_UU: "Takewhile P$UU =UU"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   130
by (simp add: Takewhile_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   131
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   132
lemma Takewhile_nil: "Takewhile P$nil =nil"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   133
by (simp add: Takewhile_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   134
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   135
lemma Takewhile_cons:
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   136
  "Takewhile P$(x\<leadsto>xs)= (if P x then x\<leadsto>(Takewhile P$xs) else nil)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   137
by (simp add: Takewhile_def Consq_def flift2_def If_and_if)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   138
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   139
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
   140
subsubsection \<open>DropWhile\<close>
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   141
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   142
lemma Dropwhile_UU: "Dropwhile P$UU =UU"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   143
by (simp add: Dropwhile_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   144
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   145
lemma Dropwhile_nil: "Dropwhile P$nil =nil"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   146
by (simp add: Dropwhile_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   147
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   148
lemma Dropwhile_cons:
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   149
  "Dropwhile P$(x\<leadsto>xs)= (if P x then Dropwhile P$xs else x\<leadsto>xs)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   150
by (simp add: Dropwhile_def Consq_def flift2_def If_and_if)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   151
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   152
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
   153
subsubsection \<open>Last\<close>
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   154
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   155
lemma Last_UU: "Last$UU =UU"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   156
by (simp add: Last_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   157
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   158
lemma Last_nil: "Last$nil =UU"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   159
by (simp add: Last_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   160
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   161
lemma Last_cons: "Last$(x\<leadsto>xs)= (if xs=nil then Def x else Last$xs)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   162
apply (simp add: Last_def Consq_def)
35532
60647586b173 adapt to changed variable name in casedist theorem
huffman
parents: 35494
diff changeset
   163
apply (cases xs)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   164
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   165
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   166
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   167
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
   168
subsubsection \<open>Flat\<close>
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   169
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   170
lemma Flat_UU: "Flat$UU =UU"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   171
by (simp add: Flat_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   172
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   173
lemma Flat_nil: "Flat$nil =nil"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   174
by (simp add: Flat_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   175
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   176
lemma Flat_cons: "Flat$(x##xs)= x @@ (Flat$xs)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   177
by (simp add: Flat_def Consq_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   178
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   179
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
   180
subsubsection \<open>Zip\<close>
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   181
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   182
lemma Zip_unfold:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   183
"Zip = (LAM t1 t2. case t1 of
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   184
                nil   => nil
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   185
              | x##xs => (case t2 of
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   186
                           nil => UU
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   187
                         | y##ys => (case x of
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   188
                                       UU  => UU
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   189
                                     | Def a => (case y of
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   190
                                                   UU => UU
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   191
                                                 | Def b => Def (a,b)##(Zip$xs$ys)))))"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   192
apply (rule trans)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   193
apply (rule fix_eq2)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   194
apply (rule Zip_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   195
apply (rule beta_cfun)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   196
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   197
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   198
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   199
lemma Zip_UU1: "Zip$UU$y =UU"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   200
apply (subst Zip_unfold)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   201
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   202
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   203
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   204
lemma Zip_UU2: "x~=nil ==> Zip$x$UU =UU"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   205
apply (subst Zip_unfold)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   206
apply simp
35532
60647586b173 adapt to changed variable name in casedist theorem
huffman
parents: 35494
diff changeset
   207
apply (cases x)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   208
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   209
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   210
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   211
lemma Zip_nil: "Zip$nil$y =nil"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   212
apply (subst Zip_unfold)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   213
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   214
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   215
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   216
lemma Zip_cons_nil: "Zip$(x\<leadsto>xs)$nil= UU"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   217
apply (subst Zip_unfold)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   218
apply (simp add: Consq_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   219
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   220
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   221
lemma Zip_cons: "Zip$(x\<leadsto>xs)$(y\<leadsto>ys)= (x,y) \<leadsto> Zip$xs$ys"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   222
apply (rule trans)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   223
apply (subst Zip_unfold)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   224
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   225
apply (simp add: Consq_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   226
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   227
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   228
lemmas [simp del] =
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   229
  sfilter_UU sfilter_nil sfilter_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   230
  smap_UU smap_nil smap_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   231
  sforall2_UU sforall2_nil sforall2_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   232
  slast_UU slast_nil slast_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   233
  stakewhile_UU  stakewhile_nil  stakewhile_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   234
  sdropwhile_UU  sdropwhile_nil  sdropwhile_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   235
  sflat_UU sflat_nil sflat_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   236
  szip_UU1 szip_UU2 szip_nil szip_cons_nil szip_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   237
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   238
lemmas [simp] =
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   239
  Filter_UU Filter_nil Filter_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   240
  Map_UU Map_nil Map_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   241
  Forall_UU Forall_nil Forall_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   242
  Last_UU Last_nil Last_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   243
  Conc_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   244
  Takewhile_UU Takewhile_nil Takewhile_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   245
  Dropwhile_UU Dropwhile_nil Dropwhile_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   246
  Zip_UU1 Zip_UU2 Zip_nil Zip_cons_nil Zip_cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   247
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   248
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   249
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   250
section "Cons"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   251
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   252
lemma Consq_def2: "a\<leadsto>s = (Def a)##s"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   253
apply (simp add: Consq_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   254
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   255
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   256
lemma Seq_exhaust: "x = UU | x = nil | (? a s. x = a \<leadsto> s)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   257
apply (simp add: Consq_def2)
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35642
diff changeset
   258
apply (cut_tac seq.nchotomy)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   259
apply (fast dest: not_Undef_is_Def [THEN iffD1])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   260
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   261
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   262
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   263
lemma Seq_cases:
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   264
"!!P. [| x = UU ==> P; x = nil ==> P; !!a s. x = a \<leadsto> s  ==> P |] ==> P"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   265
apply (cut_tac x="x" in Seq_exhaust)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   266
apply (erule disjE)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   267
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   268
apply (erule disjE)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   269
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   270
apply (erule exE)+
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   271
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   272
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   273
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   274
(*
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   275
fun Seq_case_tac s i = rule_tac x",s)] Seq_cases i
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   276
          THEN hyp_subst_tac i THEN hyp_subst_tac (i+1) THEN hyp_subst_tac (i+2);
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   277
*)
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   278
(* on a\<leadsto>s only simp_tac, as full_simp_tac is uncomplete and often causes errors *)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   279
(*
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   280
fun Seq_case_simp_tac s i = Seq_case_tac s i THEN Asm_simp_tac (i+2)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   281
                                             THEN Asm_full_simp_tac (i+1)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   282
                                             THEN Asm_full_simp_tac i;
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   283
*)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   284
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   285
lemma Cons_not_UU: "a\<leadsto>s ~= UU"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   286
apply (subst Consq_def2)
30913
10b26965a08f domain package now generates iff rules for definedness of constructors
huffman
parents: 30807
diff changeset
   287
apply simp
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   288
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   289
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   290
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   291
lemma Cons_not_less_UU: "~(a\<leadsto>x) << UU"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   292
apply (rule notI)
40431
682d6c455670 discontinue a bunch of legacy theorem names
huffman
parents: 40327
diff changeset
   293
apply (drule below_antisym)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   294
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   295
apply (simp add: Cons_not_UU)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   296
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   297
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   298
lemma Cons_not_less_nil: "~a\<leadsto>s << nil"
35117
eeec2a320a77 change generated lemmas dist_eqs and dist_les to iff-style
huffman
parents: 32149
diff changeset
   299
apply (simp add: Consq_def2)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   300
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   301
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   302
lemma Cons_not_nil: "a\<leadsto>s ~= nil"
35117
eeec2a320a77 change generated lemmas dist_eqs and dist_les to iff-style
huffman
parents: 32149
diff changeset
   303
apply (simp add: Consq_def2)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   304
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   305
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   306
lemma Cons_not_nil2: "nil ~= a\<leadsto>s"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   307
apply (simp add: Consq_def2)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   308
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   309
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   310
lemma Cons_inject_eq: "(a\<leadsto>s = b\<leadsto>t) = (a = b & s = t)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   311
apply (simp only: Consq_def2)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   312
apply (simp add: scons_inject_eq)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   313
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   314
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   315
lemma Cons_inject_less_eq: "(a\<leadsto>s<<b\<leadsto>t) = (a = b & s<<t)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   316
apply (simp add: Consq_def2)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   317
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   318
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   319
lemma seq_take_Cons: "seq_take (Suc n)$(a\<leadsto>x) = a\<leadsto> (seq_take n$x)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   320
apply (simp add: Consq_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   321
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   322
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   323
lemmas [simp] =
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   324
  Cons_not_nil2 Cons_inject_eq Cons_inject_less_eq seq_take_Cons
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   325
  Cons_not_UU Cons_not_less_UU Cons_not_less_nil Cons_not_nil
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   326
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   327
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   328
subsection "induction"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   329
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   330
lemma Seq_induct:
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   331
"!! P. [| adm P; P UU; P nil; !! a s. P s ==> P (a\<leadsto>s)|] ==> P x"
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35642
diff changeset
   332
apply (erule (2) seq.induct)
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 27239
diff changeset
   333
apply defined
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   334
apply (simp add: Consq_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   335
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   336
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   337
lemma Seq_FinitePartial_ind:
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   338
"!! P.[|P UU;P nil; !! a s. P s ==> P(a\<leadsto>s) |]
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   339
                ==> seq_finite x --> P x"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   340
apply (erule (1) seq_finite_ind)
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 27239
diff changeset
   341
apply defined
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   342
apply (simp add: Consq_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   343
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   344
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   345
lemma Seq_Finite_ind:
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   346
"!! P.[| Finite x; P nil; !! a s. [| Finite s; P s|] ==> P (a\<leadsto>s) |] ==> P x"
23778
18f426a137a9 Adapted to new inductive definition package.
berghofe
parents: 19741
diff changeset
   347
apply (erule (1) Finite.induct)
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 27239
diff changeset
   348
apply defined
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   349
apply (simp add: Consq_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   350
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   351
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   352
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   353
(* rws are definitions to be unfolded for admissibility check *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   354
(*
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   355
fun Seq_induct_tac s rws i = rule_tac x",s)] Seq_induct i
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   356
                         THEN (REPEAT_DETERM (CHANGED (Asm_simp_tac (i+1))))
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   357
                         THEN simp add: rws) i;
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   358
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   359
fun Seq_Finite_induct_tac i = erule Seq_Finite_ind i
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   360
                              THEN (REPEAT_DETERM (CHANGED (Asm_simp_tac i)));
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   361
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61032
diff changeset
   362
fun pair_tac s = rule_tac p",s)] prod.exhaust
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   363
                          THEN' hyp_subst_tac THEN' Simp_tac;
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   364
*)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   365
(* induction on a sequence of pairs with pairsplitting and simplification *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   366
(*
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   367
fun pair_induct_tac s rws i =
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   368
           rule_tac x",s)] Seq_induct i
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   369
           THEN pair_tac "a" (i+3)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   370
           THEN (REPEAT_DETERM (CHANGED (Simp_tac (i+1))))
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   371
           THEN simp add: rws) i;
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   372
*)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   373
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   374
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   375
(* ------------------------------------------------------------------------------------ *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   376
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   377
subsection "HD,TL"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   378
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   379
lemma HD_Cons [simp]: "HD$(x\<leadsto>y) = Def x"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   380
apply (simp add: Consq_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   381
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   382
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   383
lemma TL_Cons [simp]: "TL$(x\<leadsto>y) = y"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   384
apply (simp add: Consq_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   385
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   386
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   387
(* ------------------------------------------------------------------------------------ *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   388
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   389
subsection "Finite, Partial, Infinite"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   390
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   391
lemma Finite_Cons [simp]: "Finite (a\<leadsto>xs) = Finite xs"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   392
apply (simp add: Consq_def2 Finite_cons)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   393
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   394
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   395
lemma FiniteConc_1: "Finite (x::'a Seq) ==> Finite y --> Finite (x@@y)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   396
apply (erule Seq_Finite_ind, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   397
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   398
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   399
lemma FiniteConc_2:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   400
"Finite (z::'a Seq) ==> !x y. z= x@@y --> (Finite x & Finite y)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   401
apply (erule Seq_Finite_ind)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   402
(* nil*)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   403
apply (intro strip)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   404
apply (rule_tac x="x" in Seq_cases, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   405
(* cons *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   406
apply (intro strip)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   407
apply (rule_tac x="x" in Seq_cases, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   408
apply (rule_tac x="y" in Seq_cases, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   409
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   410
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   411
lemma FiniteConc [simp]: "Finite(x@@y) = (Finite (x::'a Seq) & Finite y)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   412
apply (rule iffI)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   413
apply (erule FiniteConc_2 [rule_format])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   414
apply (rule refl)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   415
apply (rule FiniteConc_1 [rule_format])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   416
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   417
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   418
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   419
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   420
lemma FiniteMap1: "Finite s ==> Finite (Map f$s)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   421
apply (erule Seq_Finite_ind, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   422
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   423
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   424
lemma FiniteMap2: "Finite s ==> ! t. (s = Map f$t) --> Finite t"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   425
apply (erule Seq_Finite_ind)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   426
apply (intro strip)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   427
apply (rule_tac x="t" in Seq_cases, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   428
(* main case *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   429
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   430
apply (rule_tac x="t" in Seq_cases, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   431
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   432
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   433
lemma Map2Finite: "Finite (Map f$s) = Finite s"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   434
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   435
apply (erule FiniteMap2 [rule_format])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   436
apply (rule refl)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   437
apply (erule FiniteMap1)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   438
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   439
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   440
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   441
lemma FiniteFilter: "Finite s ==> Finite (Filter P$s)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   442
apply (erule Seq_Finite_ind, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   443
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   444
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   445
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   446
(* ----------------------------------------------------------------------------------- *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   447
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   448
subsection "Conc"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   449
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   450
lemma Conc_cong: "!! x::'a Seq. Finite x ==> ((x @@ y) = (x @@ z)) = (y = z)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   451
apply (erule Seq_Finite_ind, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   452
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   453
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   454
lemma Conc_assoc: "(x @@ y) @@ z = (x::'a Seq) @@ y @@ z"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   455
apply (rule_tac x="x" in Seq_induct, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   456
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   457
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   458
lemma nilConc [simp]: "s@@ nil = s"
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35642
diff changeset
   459
apply (induct s)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   460
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   461
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   462
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   463
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   464
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   465
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   466
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   467
(* should be same as nil_is_Conc2 when all nils are turned to right side !! *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   468
lemma nil_is_Conc: "(nil = x @@ y) = ((x::'a Seq)= nil & y = nil)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   469
apply (rule_tac x="x" in Seq_cases)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   470
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   471
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   472
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   473
lemma nil_is_Conc2: "(x @@ y = nil) = ((x::'a Seq)= nil & y = nil)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   474
apply (rule_tac x="x" in Seq_cases)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   475
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   476
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   477
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   478
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   479
(* ------------------------------------------------------------------------------------ *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   480
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   481
subsection "Last"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   482
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   483
lemma Finite_Last1: "Finite s ==> s~=nil --> Last$s~=UU"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   484
apply (erule Seq_Finite_ind, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   485
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   486
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   487
lemma Finite_Last2: "Finite s ==> Last$s=UU --> s=nil"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   488
apply (erule Seq_Finite_ind, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   489
apply fast
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   490
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   491
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   492
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   493
(* ------------------------------------------------------------------------------------ *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   494
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   495
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   496
subsection "Filter, Conc"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   497
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   498
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   499
lemma FilterPQ: "Filter P$(Filter Q$s) = Filter (%x. P x & Q x)$s"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   500
apply (rule_tac x="s" in Seq_induct, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   501
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   502
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   503
lemma FilterConc: "Filter P$(x @@ y) = (Filter P$x @@ Filter P$y)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   504
apply (simp add: Filter_def sfiltersconc)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   505
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   506
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   507
(* ------------------------------------------------------------------------------------ *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   508
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   509
subsection "Map"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   510
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   511
lemma MapMap: "Map f$(Map g$s) = Map (f o g)$s"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   512
apply (rule_tac x="s" in Seq_induct, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   513
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   514
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   515
lemma MapConc: "Map f$(x@@y) = (Map f$x) @@ (Map f$y)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   516
apply (rule_tac x="x" in Seq_induct, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   517
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   518
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   519
lemma MapFilter: "Filter P$(Map f$x) = Map f$(Filter (P o f)$x)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   520
apply (rule_tac x="x" in Seq_induct, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   521
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   522
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   523
lemma nilMap: "nil = (Map f$s) --> s= nil"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   524
apply (rule_tac x="s" in Seq_cases, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   525
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   526
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   527
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   528
lemma ForallMap: "Forall P (Map f$s) = Forall (P o f) s"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   529
apply (rule_tac x="s" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   530
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   531
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   532
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   533
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   534
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   535
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   536
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   537
(* ------------------------------------------------------------------------------------ *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   538
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   539
subsection "Forall"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   540
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   541
26008
24c82bef5696 eliminated escaped white space;
wenzelm
parents: 25923
diff changeset
   542
lemma ForallPForallQ1: "Forall P ys & (! x. P x --> Q x)
24c82bef5696 eliminated escaped white space;
wenzelm
parents: 25923
diff changeset
   543
         --> Forall Q ys"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   544
apply (rule_tac x="ys" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   545
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   546
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   547
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   548
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   549
lemmas ForallPForallQ =
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   550
  ForallPForallQ1 [THEN mp, OF conjI, OF _ allI, OF _ impI]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   552
lemma Forall_Conc_impl: "(Forall P x & Forall P y) --> Forall P (x @@ y)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   553
apply (rule_tac x="x" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   554
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   555
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   556
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   557
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   558
lemma Forall_Conc [simp]:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   559
  "Finite x ==> Forall P (x @@ y) = (Forall P x & Forall P y)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   560
apply (erule Seq_Finite_ind, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   561
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   562
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   563
lemma ForallTL1: "Forall P s  --> Forall P (TL$s)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   564
apply (rule_tac x="s" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   565
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   566
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   567
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   568
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   569
lemmas ForallTL = ForallTL1 [THEN mp]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   570
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   571
lemma ForallDropwhile1: "Forall P s  --> Forall P (Dropwhile Q$s)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   572
apply (rule_tac x="s" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   573
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   574
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   575
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   576
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   577
lemmas ForallDropwhile = ForallDropwhile1 [THEN mp]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   578
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   579
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   580
(* only admissible in t, not if done in s *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   581
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   582
lemma Forall_prefix: "! s. Forall P s --> t<<s --> Forall P t"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   583
apply (rule_tac x="t" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   584
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   585
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   586
apply (intro strip)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   587
apply (rule_tac x="sa" in Seq_cases)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   588
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   589
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   590
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   591
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   592
lemmas Forall_prefixclosed = Forall_prefix [rule_format]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   593
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   594
lemma Forall_postfixclosed:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   595
  "[| Finite h; Forall P s; s= h @@ t |] ==> Forall P t"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   596
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   597
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   598
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   599
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   600
lemma ForallPFilterQR1:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   601
  "((! x. P x --> (Q x = R x)) & Forall P tr) --> Filter Q$tr = Filter R$tr"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   602
apply (rule_tac x="tr" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   603
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   604
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   605
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   606
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   607
lemmas ForallPFilterQR = ForallPFilterQR1 [THEN mp, OF conjI, OF allI]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   608
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   609
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   610
(* ------------------------------------------------------------------------------------- *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   611
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   612
subsection "Forall, Filter"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   613
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   614
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   615
lemma ForallPFilterP: "Forall P (Filter P$x)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   616
apply (simp add: Filter_def Forall_def forallPsfilterP)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   617
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   618
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   619
(* holds also in other direction, then equal to forallPfilterP *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   620
lemma ForallPFilterPid1: "Forall P x --> Filter P$x = x"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   621
apply (rule_tac x="x" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   622
apply (simp add: Forall_def sforall_def Filter_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   623
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   624
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   625
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   626
lemmas ForallPFilterPid = ForallPFilterPid1 [THEN mp]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   627
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   628
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   629
(* holds also in other direction *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   630
lemma ForallnPFilterPnil1: "!! ys . Finite ys ==>
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   631
   Forall (%x. ~P x) ys --> Filter P$ys = nil "
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   632
apply (erule Seq_Finite_ind, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   633
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   634
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   635
lemmas ForallnPFilterPnil = ForallnPFilterPnil1 [THEN mp]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   636
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   637
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   638
(* holds also in other direction *)
26008
24c82bef5696 eliminated escaped white space;
wenzelm
parents: 25923
diff changeset
   639
lemma ForallnPFilterPUU1: "~Finite ys & Forall (%x. ~P x) ys
24c82bef5696 eliminated escaped white space;
wenzelm
parents: 25923
diff changeset
   640
                  --> Filter P$ys = UU "
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   641
apply (rule_tac x="ys" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   642
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   643
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   644
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   645
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   646
lemmas ForallnPFilterPUU = ForallnPFilterPUU1 [THEN mp, OF conjI]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   647
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   648
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   649
(* inverse of ForallnPFilterPnil *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   650
48194
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   651
lemma FilternPnilForallP [rule_format]: "Filter P$ys = nil -->
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   652
   (Forall (%x. ~P x) ys & Finite ys)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   653
apply (rule_tac x="ys" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   654
(* adm *)
30807
a167ed35ec0d domain package declares more simp rules
huffman
parents: 30607
diff changeset
   655
apply (simp add: Forall_def sforall_def)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   656
(* base cases *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   657
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   658
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   659
(* main case *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   660
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   661
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   662
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   663
48194
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   664
(* inverse of ForallnPFilterPUU *)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   665
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   666
lemma FilternPUUForallP:
48194
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   667
  assumes "Filter P$ys = UU"
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   668
  shows "Forall (%x. ~P x) ys  & ~Finite ys"
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   669
proof
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   670
  show "Forall (%x. ~P x) ys"
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   671
  proof (rule classical)
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   672
    assume "\<not> ?thesis"
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   673
    then have "Filter P$ys ~= UU"
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   674
      apply (rule rev_mp)
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   675
      apply (induct ys rule: Seq_induct)
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   676
      apply (simp add: Forall_def sforall_def)
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   677
      apply simp_all
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   678
      done
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   679
    with assms show ?thesis by contradiction
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   680
  qed
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   681
  show "~ Finite ys"
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   682
  proof
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   683
    assume "Finite ys"
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   684
    then have "Filter P$ys ~= UU"
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   685
      by (rule Seq_Finite_ind) simp_all
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   686
    with assms show False by contradiction
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   687
  qed
1440a3103af0 tuned proofs -- eliminated old-fashioned COMP and rev_contrapos;
wenzelm
parents: 44890
diff changeset
   688
qed
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   689
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   690
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   691
lemma ForallQFilterPnil:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   692
  "!! Q P.[| Forall Q ys; Finite ys; !!x. Q x ==> ~P x|]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   693
    ==> Filter P$ys = nil"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   694
apply (erule ForallnPFilterPnil)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   695
apply (erule ForallPForallQ)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   696
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   697
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   698
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   699
lemma ForallQFilterPUU:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   700
 "!! Q P. [| ~Finite ys; Forall Q ys;  !!x. Q x ==> ~P x|]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   701
    ==> Filter P$ys = UU "
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   702
apply (erule ForallnPFilterPUU)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   703
apply (erule ForallPForallQ)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   704
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   705
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   706
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   707
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   708
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   709
(* ------------------------------------------------------------------------------------- *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   710
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   711
subsection "Takewhile, Forall, Filter"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   712
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   713
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   714
lemma ForallPTakewhileP [simp]: "Forall P (Takewhile P$x)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   715
apply (simp add: Forall_def Takewhile_def sforallPstakewhileP)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   716
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   717
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   718
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   719
lemma ForallPTakewhileQ [simp]:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   720
"!! P. [| !!x. Q x==> P x |] ==> Forall P (Takewhile Q$x)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   721
apply (rule ForallPForallQ)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   722
apply (rule ForallPTakewhileP)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   723
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   724
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   725
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   726
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   727
lemma FilterPTakewhileQnil [simp]:
26008
24c82bef5696 eliminated escaped white space;
wenzelm
parents: 25923
diff changeset
   728
  "!! Q P.[| Finite (Takewhile Q$ys); !!x. Q x ==> ~P x |]
24c82bef5696 eliminated escaped white space;
wenzelm
parents: 25923
diff changeset
   729
   ==> Filter P$(Takewhile Q$ys) = nil"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   730
apply (erule ForallnPFilterPnil)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   731
apply (rule ForallPForallQ)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   732
apply (rule ForallPTakewhileP)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   733
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   734
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   735
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   736
lemma FilterPTakewhileQid [simp]:
26008
24c82bef5696 eliminated escaped white space;
wenzelm
parents: 25923
diff changeset
   737
 "!! Q P. [| !!x. Q x ==> P x |] ==>
24c82bef5696 eliminated escaped white space;
wenzelm
parents: 25923
diff changeset
   738
            Filter P$(Takewhile Q$ys) = (Takewhile Q$ys)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   739
apply (rule ForallPFilterPid)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   740
apply (rule ForallPForallQ)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   741
apply (rule ForallPTakewhileP)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   742
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   743
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   744
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   745
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   746
lemma Takewhile_idempotent: "Takewhile P$(Takewhile P$s) = Takewhile P$s"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   747
apply (rule_tac x="s" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   748
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   749
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   750
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   751
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   752
lemma ForallPTakewhileQnP [simp]:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   753
 "Forall P s --> Takewhile (%x. Q x | (~P x))$s = Takewhile Q$s"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   754
apply (rule_tac x="s" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   755
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   756
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   757
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   758
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   759
lemma ForallPDropwhileQnP [simp]:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   760
 "Forall P s --> Dropwhile (%x. Q x | (~P x))$s = Dropwhile Q$s"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   761
apply (rule_tac x="s" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   762
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   763
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   764
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   765
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   766
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   767
lemma TakewhileConc1:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   768
 "Forall P s --> Takewhile P$(s @@ t) = s @@ (Takewhile P$t)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   769
apply (rule_tac x="s" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   770
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   771
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   772
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   773
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   774
lemmas TakewhileConc = TakewhileConc1 [THEN mp]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   775
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   776
lemma DropwhileConc1:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   777
 "Finite s ==> Forall P s --> Dropwhile P$(s @@ t) = Dropwhile P$t"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   778
apply (erule Seq_Finite_ind, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   779
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   780
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   781
lemmas DropwhileConc = DropwhileConc1 [THEN mp]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   782
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   783
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   784
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   785
(* ----------------------------------------------------------------------------------- *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   786
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   787
subsection "coinductive characterizations of Filter"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   788
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   789
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   790
lemma divide_Seq_lemma:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   791
 "HD$(Filter P$y) = Def x
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   792
    --> y = ((Takewhile (%x. ~P x)$y) @@ (x \<leadsto> TL$(Dropwhile (%a. ~P a)$y)))
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   793
             & Finite (Takewhile (%x. ~ P x)$y)  & P x"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   794
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   795
(* FIX: pay attention: is only admissible with chain-finite package to be added to
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   796
        adm test and Finite f x admissibility *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   797
apply (rule_tac x="y" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   798
apply (simp add: adm_subst [OF _ adm_Finite])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   799
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   800
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   801
apply (case_tac "P a")
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   802
 apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   803
 apply blast
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   804
(* ~ P a *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   805
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   806
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   807
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   808
lemma divide_Seq: "(x\<leadsto>xs) << Filter P$y 
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   809
   ==> y = ((Takewhile (%a. ~ P a)$y) @@ (x \<leadsto> TL$(Dropwhile (%a. ~ P a)$y)))
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   810
      & Finite (Takewhile (%a. ~ P a)$y)  & P x"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   811
apply (rule divide_Seq_lemma [THEN mp])
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   812
apply (drule_tac f="HD" and x="x\<leadsto>xs" in  monofun_cfun_arg)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   813
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   814
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   815
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   816
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   817
lemma nForall_HDFilter:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   818
 "~Forall P y --> (? x. HD$(Filter (%a. ~P a)$y) = Def x)"
35907
ea0bf2a01eb0 avoid dependence on adm_tac solver
huffman
parents: 35781
diff changeset
   819
unfolding not_Undef_is_Def [symmetric]
ea0bf2a01eb0 avoid dependence on adm_tac solver
huffman
parents: 35781
diff changeset
   820
apply (induct y rule: Seq_induct)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   821
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   822
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   823
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   824
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   825
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   826
lemma divide_Seq2: "~Forall P y
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   827
  ==> ? x. y= (Takewhile P$y @@ (x \<leadsto> TL$(Dropwhile P$y))) &
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   828
      Finite (Takewhile P$y) & (~ P x)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   829
apply (drule nForall_HDFilter [THEN mp])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   830
apply safe
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   831
apply (rule_tac x="x" in exI)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   832
apply (cut_tac P1="%x. ~ P x" in divide_Seq_lemma [THEN mp])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   833
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   834
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   835
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   836
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   837
lemma divide_Seq3: "~Forall P y
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   838
  ==> ? x bs rs. y= (bs @@ (x\<leadsto>rs)) & Finite bs & Forall P bs & (~ P x)"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   839
apply (drule divide_Seq2)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   840
(*Auto_tac no longer proves it*)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42151
diff changeset
   841
apply fastforce
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   842
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   843
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   844
lemmas [simp] = FilterPQ FilterConc Conc_cong
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   845
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   846
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   847
(* ------------------------------------------------------------------------------------- *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   848
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   849
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   850
subsection "take_lemma"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   851
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   852
lemma seq_take_lemma: "(!n. seq_take n$x = seq_take n$x') = (x = x')"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   853
apply (rule iffI)
35642
f478d5a9d238 generate separate qualified theorem name for each type's reach and take_lemma
huffman
parents: 35532
diff changeset
   854
apply (rule seq.take_lemma)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   855
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   856
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   857
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   858
lemma take_reduction1:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   859
"  ! n. ((! k. k < n --> seq_take k$y1 = seq_take k$y2)
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   860
    --> seq_take n$(x @@ (t\<leadsto>y1)) =  seq_take n$(x @@ (t\<leadsto>y2)))"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   861
apply (rule_tac x="x" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   862
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   863
apply (intro strip)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   864
apply (case_tac "n")
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   865
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   866
apply (case_tac "n")
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   867
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   868
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   869
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   870
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   871
lemma take_reduction:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   872
 "!! n.[| x=y; s=t; !! k. k<n ==> seq_take k$y1 = seq_take k$y2|]
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   873
  ==> seq_take n$(x @@ (s\<leadsto>y1)) =  seq_take n$(y @@ (t\<leadsto>y2))"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   874
apply (auto intro!: take_reduction1 [rule_format])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   875
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   876
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   877
(* ------------------------------------------------------------------
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   878
          take-lemma and take_reduction for << instead of =
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   879
   ------------------------------------------------------------------ *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   880
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   881
lemma take_reduction_less1:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   882
"  ! n. ((! k. k < n --> seq_take k$y1 << seq_take k$y2)
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   883
    --> seq_take n$(x @@ (t\<leadsto>y1)) <<  seq_take n$(x @@ (t\<leadsto>y2)))"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   884
apply (rule_tac x="x" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   885
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   886
apply (intro strip)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   887
apply (case_tac "n")
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   888
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   889
apply (case_tac "n")
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   890
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   891
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   892
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   893
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   894
lemma take_reduction_less:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   895
 "!! n.[| x=y; s=t;!! k. k<n ==> seq_take k$y1 << seq_take k$y2|]
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   896
  ==> seq_take n$(x @@ (s\<leadsto>y1)) <<  seq_take n$(y @@ (t\<leadsto>y2))"
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   897
apply (auto intro!: take_reduction_less1 [rule_format])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   898
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   899
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   900
lemma take_lemma_less1:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   901
  assumes "!! n. seq_take n$s1 << seq_take n$s2"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   902
  shows "s1<<s2"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   903
apply (rule_tac t="s1" in seq.reach [THEN subst])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   904
apply (rule_tac t="s2" in seq.reach [THEN subst])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   905
apply (rule lub_mono)
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 37140
diff changeset
   906
apply (rule seq.chain_take [THEN ch2ch_Rep_cfunL])
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 37140
diff changeset
   907
apply (rule seq.chain_take [THEN ch2ch_Rep_cfunL])
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35259
diff changeset
   908
apply (rule assms)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   909
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   910
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   911
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   912
lemma take_lemma_less: "(!n. seq_take n$x << seq_take n$x') = (x << x')"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   913
apply (rule iffI)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   914
apply (rule take_lemma_less1)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   915
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   916
apply (erule monofun_cfun_arg)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   917
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   918
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   919
(* ------------------------------------------------------------------
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   920
          take-lemma proof principles
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   921
   ------------------------------------------------------------------ *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   922
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   923
lemma take_lemma_principle1:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   924
 "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   925
            !! s1 s2 y. [| Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2)|]
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   926
                          ==> (f (s1 @@ y\<leadsto>s2)) = (g (s1 @@ y\<leadsto>s2)) |]
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   927
               ==> A x --> (f x)=(g x)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   928
apply (case_tac "Forall Q x")
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   929
apply (auto dest!: divide_Seq3)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   930
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   931
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   932
lemma take_lemma_principle2:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   933
  "!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   934
           !! s1 s2 y. [| Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2)|]
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   935
                          ==> ! n. seq_take n$(f (s1 @@ y\<leadsto>s2))
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   936
                                 = seq_take n$(g (s1 @@ y\<leadsto>s2)) |]
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   937
               ==> A x --> (f x)=(g x)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   938
apply (case_tac "Forall Q x")
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   939
apply (auto dest!: divide_Seq3)
35642
f478d5a9d238 generate separate qualified theorem name for each type's reach and take_lemma
huffman
parents: 35532
diff changeset
   940
apply (rule seq.take_lemma)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   941
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   942
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   943
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   944
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   945
(* Note: in the following proofs the ordering of proof steps is very
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   946
         important, as otherwise either (Forall Q s1) would be in the IH as
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   947
         assumption (then rule useless) or it is not possible to strengthen
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   948
         the IH apply doing a forall closure of the sequence t (then rule also useless).
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   949
         This is also the reason why the induction rule (nat_less_induct or nat_induct) has to
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   950
         to be imbuilt into the rule, as induction has to be done early and the take lemma
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   951
         has to be used in the trivial direction afterwards for the (Forall Q x) case.  *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   952
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   953
lemma take_lemma_induct:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   954
"!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   955
         !! s1 s2 y n. [| ! t. A t --> seq_take n$(f t) = seq_take n$(g t);
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   956
                          Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2) |]
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   957
                          ==>   seq_take (Suc n)$(f (s1 @@ y\<leadsto>s2))
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   958
                              = seq_take (Suc n)$(g (s1 @@ y\<leadsto>s2)) |]
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   959
               ==> A x --> (f x)=(g x)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   960
apply (rule impI)
35642
f478d5a9d238 generate separate qualified theorem name for each type's reach and take_lemma
huffman
parents: 35532
diff changeset
   961
apply (rule seq.take_lemma)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   962
apply (rule mp)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   963
prefer 2 apply assumption
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   964
apply (rule_tac x="x" in spec)
27105
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 26008
diff changeset
   965
apply (rule nat.induct)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   966
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   967
apply (rule allI)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   968
apply (case_tac "Forall Q xa")
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   969
apply (force intro!: seq_take_lemma [THEN iffD2, THEN spec])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   970
apply (auto dest!: divide_Seq3)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   971
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   972
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   973
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   974
lemma take_lemma_less_induct:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   975
"!! Q. [|!! s. [| Forall Q s; A s |] ==> (f s) = (g s) ;
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   976
        !! s1 s2 y n. [| ! t m. m < n --> A t --> seq_take m$(f t) = seq_take m$(g t);
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   977
                          Forall Q s1; Finite s1; ~ Q y; A (s1 @@ y\<leadsto>s2) |]
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   978
                          ==>   seq_take n$(f (s1 @@ y\<leadsto>s2))
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   979
                              = seq_take n$(g (s1 @@ y\<leadsto>s2)) |]
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   980
               ==> A x --> (f x)=(g x)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   981
apply (rule impI)
35642
f478d5a9d238 generate separate qualified theorem name for each type's reach and take_lemma
huffman
parents: 35532
diff changeset
   982
apply (rule seq.take_lemma)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   983
apply (rule mp)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   984
prefer 2 apply assumption
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   985
apply (rule_tac x="x" in spec)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   986
apply (rule nat_less_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   987
apply (rule allI)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   988
apply (case_tac "Forall Q xa")
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   989
apply (force intro!: seq_take_lemma [THEN iffD2, THEN spec])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   990
apply (auto dest!: divide_Seq3)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   991
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   992
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   993
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   994
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   995
lemma take_lemma_in_eq_out:
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   996
"!! Q. [| A UU  ==> (f UU) = (g UU) ;
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   997
          A nil ==> (f nil) = (g nil) ;
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
   998
          !! s y n. [| ! t. A t --> seq_take n$(f t) = seq_take n$(g t);
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
   999
                     A (y\<leadsto>s) |]
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
  1000
                     ==>   seq_take (Suc n)$(f (y\<leadsto>s))
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
  1001
                         = seq_take (Suc n)$(g (y\<leadsto>s)) |]
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1002
               ==> A x --> (f x)=(g x)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1003
apply (rule impI)
35642
f478d5a9d238 generate separate qualified theorem name for each type's reach and take_lemma
huffman
parents: 35532
diff changeset
  1004
apply (rule seq.take_lemma)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1005
apply (rule mp)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1006
prefer 2 apply assumption
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1007
apply (rule_tac x="x" in spec)
27105
5f139027c365 slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
haftmann
parents: 26008
diff changeset
  1008
apply (rule nat.induct)
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1009
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1010
apply (rule allI)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1011
apply (rule_tac x="xa" in Seq_cases)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1012
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1013
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1014
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1015
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1016
(* ------------------------------------------------------------------------------------ *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1017
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1018
subsection "alternative take_lemma proofs"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1019
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1020
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1021
(* --------------------------------------------------------------- *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1022
(*              Alternative Proof of FilterPQ                      *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1023
(* --------------------------------------------------------------- *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1024
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1025
declare FilterPQ [simp del]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1026
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1027
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1028
(* In general: How to do this case without the same adm problems
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1029
   as for the entire proof ? *)
62000
8cba509ace9c more symbols;
wenzelm
parents: 61424
diff changeset
  1030
lemma Filter_lemma1: "Forall (%x. ~(P x & Q x)) s
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1031
          --> Filter P$(Filter Q$s) =
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1032
              Filter (%x. P x & Q x)$s"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1033
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1034
apply (rule_tac x="s" in Seq_induct)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1035
apply (simp add: Forall_def sforall_def)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1036
apply simp_all
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1037
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1038
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1039
lemma Filter_lemma2: "Finite s ==>
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1040
          (Forall (%x. (~P x) | (~ Q x)) s
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1041
          --> Filter P$(Filter Q$s) = nil)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1042
apply (erule Seq_Finite_ind, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1043
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1044
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1045
lemma Filter_lemma3: "Finite s ==>
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1046
          Forall (%x. (~P x) | (~ Q x)) s
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1047
          --> Filter (%x. P x & Q x)$s = nil"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1048
apply (erule Seq_Finite_ind, simp_all)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1049
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1050
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1051
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1052
lemma FilterPQ_takelemma: "Filter P$(Filter Q$s) = Filter (%x. P x & Q x)$s"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1053
apply (rule_tac A1="%x. True" and
62000
8cba509ace9c more symbols;
wenzelm
parents: 61424
diff changeset
  1054
                 Q1="%x. ~(P x & Q x)" and x1="s" in
19551
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1055
                 take_lemma_induct [THEN mp])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1056
(* better support for A = %x. True *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1057
apply (simp add: Filter_lemma1)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1058
apply (simp add: Filter_lemma2 Filter_lemma3)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1059
apply simp
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1060
done
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1061
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1062
declare FilterPQ [simp]
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1063
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1064
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1065
(* --------------------------------------------------------------- *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1066
(*              Alternative Proof of MapConc                       *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1067
(* --------------------------------------------------------------- *)
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1068
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1069
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1070
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1071
lemma MapConc_takelemma: "Map f$(x@@y) = (Map f$x) @@ (Map f$y)"
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1072
apply (rule_tac A1="%x. True" and x1="x" in
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1073
    take_lemma_in_eq_out [THEN mp])
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1074
apply auto
4103954f3668 converted to isar theory; removed unsound adm_all axiom
huffman
parents: 17233
diff changeset
  1075
done
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
  1076
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1077
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
  1078
ML \<open>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1079
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27105
diff changeset
  1080
fun Seq_case_tac ctxt s i =
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
  1081
  Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_cases} i
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
  1082
  THEN hyp_subst_tac ctxt i THEN hyp_subst_tac ctxt (i+1) THEN hyp_subst_tac ctxt (i+2);
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1083
62001
1f2788fb0b8b more symbols;
wenzelm
parents: 62000
diff changeset
  1084
(* on a\<leadsto>s only simp_tac, as full_simp_tac is uncomplete and often causes errors *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27105
diff changeset
  1085
fun Seq_case_simp_tac ctxt s i =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48194
diff changeset
  1086
  Seq_case_tac ctxt s i
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48194
diff changeset
  1087
  THEN asm_simp_tac ctxt (i+2)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48194
diff changeset
  1088
  THEN asm_full_simp_tac ctxt (i+1)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48194
diff changeset
  1089
  THEN asm_full_simp_tac ctxt i;
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1090
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1091
(* rws are definitions to be unfolded for admissibility check *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27105
diff changeset
  1092
fun Seq_induct_tac ctxt s rws i =
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
  1093
  Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_induct} i
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48194
diff changeset
  1094
  THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ctxt (i+1))))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48194
diff changeset
  1095
  THEN simp_tac (ctxt addsimps rws) i;
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1096
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27105
diff changeset
  1097
fun Seq_Finite_induct_tac ctxt i =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59780
diff changeset
  1098
  eresolve_tac ctxt @{thms Seq_Finite_ind} i
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48194
diff changeset
  1099
  THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ctxt i)));
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1100
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27105
diff changeset
  1101
fun pair_tac ctxt s =
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61032
diff changeset
  1102
  Rule_Insts.res_inst_tac ctxt [((("y", 0), Position.none), s)] [] @{thm prod.exhaust}
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51717
diff changeset
  1103
  THEN' hyp_subst_tac ctxt THEN' asm_full_simp_tac ctxt;
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1104
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1105
(* induction on a sequence of pairs with pairsplitting and simplification *)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27105
diff changeset
  1106
fun pair_induct_tac ctxt s rws i =
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
  1107
  Rule_Insts.res_inst_tac ctxt [((("x", 0), Position.none), s)] [] @{thm Seq_induct} i
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48194
diff changeset
  1108
  THEN pair_tac ctxt "a" (i+3)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48194
diff changeset
  1109
  THEN (REPEAT_DETERM (CHANGED (simp_tac ctxt (i+1))))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48194
diff changeset
  1110
  THEN simp_tac (ctxt addsimps rws) i;
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1111
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
  1112
\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1113
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 19551
diff changeset
  1114
end