| author | haftmann | 
| Fri, 05 Dec 2008 18:43:42 +0100 | |
| changeset 29006 | abe0f11cfa4e | 
| parent 26793 | e36a92ff543e | 
| child 30198 | 922f944f03b2 | 
| permissions | -rw-r--r-- | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
1  | 
(* Title: HOL/Induct/LList.thy  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
5  | 
Shares NIL, CONS, List_case with List.thy  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
6  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
7  | 
Still needs flatten function -- hard because it need  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
8  | 
bounds on the amount of lookahead required.  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
10  | 
Could try (but would it work for the gfp analogue of term?)  | 
| 3842 | 11  | 
  LListD_Fun_def "LListD_Fun(A) == (%Z. diag({Numb(0)}) <++> diag(A) <**> Z)"
 | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
12  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
13  | 
A nice but complex example would be [ML for the Working Programmer, page 176]  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
14  | 
from(1) = enumerate (Lmap (Lmap(pack), makeqq(from(1),from(1))))  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
15  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
16  | 
Previous definition of llistD_Fun was explicit:  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
17  | 
llistD_Fun_def  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
18  | 
"llistD_Fun(r) ==  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
19  | 
       {(LNil,LNil)}  Un        
 | 
| 10834 | 20  | 
(UN x. (split(%l1 l2.(LCons(x,l1),LCons(x,l2))))`r)"  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
21  | 
*)  | 
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
22  | 
|
| 13107 | 23  | 
header {*Definition of type llist by a greatest fixed point*}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
24  | 
|
| 20801 | 25  | 
theory LList imports SList begin  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
26  | 
|
| 23746 | 27  | 
coinductive_set  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
28  | 
llist :: "'a item set => 'a item set"  | 
| 23746 | 29  | 
for A :: "'a item set"  | 
30  | 
where  | 
|
31  | 
NIL_I: "NIL \<in> llist(A)"  | 
|
32  | 
| CONS_I: "[| a \<in> A; M \<in> llist(A) |] ==> CONS a M \<in> llist(A)"  | 
|
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
33  | 
|
| 23746 | 34  | 
coinductive_set  | 
35  | 
  LListD :: "('a item * 'a item)set => ('a item * 'a item)set"
 | 
|
36  | 
  for r :: "('a item * 'a item)set"
 | 
|
37  | 
where  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
38  | 
NIL_I: "(NIL, NIL) \<in> LListD(r)"  | 
| 23746 | 39  | 
| CONS_I: "[| (a,b) \<in> r; (M,N) \<in> LListD(r) |]  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
40  | 
==> (CONS a M, CONS b N) \<in> LListD(r)"  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
41  | 
|
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
42  | 
|
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
43  | 
typedef (LList)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
44  | 
'a llist = "llist(range Leaf) :: 'a item set"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
45  | 
by (blast intro: llist.NIL_I)  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
46  | 
|
| 19736 | 47  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
48  | 
list_Fun :: "['a item set, 'a item set] => 'a item set" where  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
49  | 
    --{*Now used exclusively for abbreviating the coinduction rule*}
 | 
| 19736 | 50  | 
     "list_Fun A X = {z. z = NIL | (\<exists>M a. z = CONS a M & a \<in> A & M \<in> X)}"
 | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
51  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
52  | 
definition  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
53  | 
LListD_Fun ::  | 
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
54  | 
      "[('a item * 'a item)set, ('a item * 'a item)set] => 
 | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
55  | 
       ('a item * 'a item)set" where
 | 
| 19736 | 56  | 
"LListD_Fun r X =  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
57  | 
       {z. z = (NIL, NIL) |   
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
58  | 
(\<exists>M N a b. z = (CONS a M, CONS b N) & (a, b) \<in> r & (M, N) \<in> X)}"  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
59  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
60  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
61  | 
LNil :: "'a llist" where  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
62  | 
     --{*abstract constructor*}
 | 
| 19736 | 63  | 
"LNil = Abs_LList NIL"  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
64  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
65  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
66  | 
LCons :: "['a, 'a llist] => 'a llist" where  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
67  | 
     --{*abstract constructor*}
 | 
| 19736 | 68  | 
"LCons x xs = Abs_LList(CONS (Leaf x) (Rep_LList xs))"  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
69  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
70  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
71  | 
llist_case :: "['b, ['a, 'a llist]=>'b, 'a llist] => 'b" where  | 
| 19736 | 72  | 
"llist_case c d l =  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
73  | 
List_case c (%x y. d (inv Leaf x) (Abs_LList y)) (Rep_LList l)"  | 
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
74  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
75  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
76  | 
  LList_corec_fun :: "[nat, 'a=> ('b item * 'a) option, 'a] => 'b item" where
 | 
| 19736 | 77  | 
"LList_corec_fun k f ==  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
78  | 
     nat_rec (%x. {})                         
 | 
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
79  | 
(%j r x. case f x of None => NIL  | 
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
80  | 
| Some(z,w) => CONS z (r w))  | 
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
81  | 
k"  | 
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
82  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
83  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
84  | 
  LList_corec     :: "['a, 'a => ('b item * 'a) option] => 'b item" where
 | 
| 19736 | 85  | 
"LList_corec a f = (\<Union>k. LList_corec_fun k f a)"  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
86  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
87  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
88  | 
  llist_corec     :: "['a, 'a => ('b * 'a) option] => 'b llist" where
 | 
| 19736 | 89  | 
"llist_corec a f =  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
90  | 
Abs_LList(LList_corec a  | 
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
91  | 
(%z. case f z of None => None  | 
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
92  | 
| Some(v,w) => Some(Leaf(v), w)))"  | 
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
93  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
94  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
95  | 
  llistD_Fun :: "('a llist * 'a llist)set => ('a llist * 'a llist)set" where
 | 
| 19736 | 96  | 
"llistD_Fun(r) =  | 
| 10834 | 97  | 
prod_fun Abs_LList Abs_LList `  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
98  | 
LListD_Fun (diag(range Leaf))  | 
| 10834 | 99  | 
(prod_fun Rep_LList Rep_LList ` r)"  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
100  | 
|
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
101  | 
|
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
102  | 
|
| 13107 | 103  | 
text{* The case syntax for type @{text "'a llist"} *}
 | 
| 20770 | 104  | 
syntax (* FIXME proper case syntax!? *)  | 
105  | 
LNil :: logic  | 
|
106  | 
LCons :: logic  | 
|
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
107  | 
translations  | 
| 20770 | 108  | 
"case p of LNil => a | LCons x l => b" == "CONST llist_case a (%x l. b) p"  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
110  | 
|
| 13107 | 111  | 
subsubsection{* Sample function definitions.  Item-based ones start with @{text L} *}
 | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
112  | 
|
| 19736 | 113  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
114  | 
  Lmap       :: "('a item => 'b item) => ('a item => 'b item)" where
 | 
| 19736 | 115  | 
"Lmap f M = LList_corec M (List_case None (%x M'. Some((f(x), M'))))"  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
116  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
117  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
118  | 
  lmap       :: "('a=>'b) => ('a llist => 'b llist)" where
 | 
| 19736 | 119  | 
"lmap f l = llist_corec l (%z. case z of LNil => None  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
120  | 
| LCons y z => Some(f(y), z))"  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
121  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
122  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
123  | 
iterates :: "['a => 'a, 'a] => 'a llist" where  | 
| 19736 | 124  | 
"iterates f a = llist_corec a (%x. Some((x, f(x))))"  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
125  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
126  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
127  | 
Lconst :: "'a item => 'a item" where  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
128  | 
"Lconst(M) == lfp(%N. CONS M N)"  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
129  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
130  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
131  | 
Lappend :: "['a item, 'a item] => 'a item" where  | 
| 19736 | 132  | 
"Lappend M N = LList_corec (M,N)  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
133  | 
(split(List_case (List_case None (%N1 N2. Some((N1, (NIL,N2)))))  | 
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
134  | 
(%M1 M2 N. Some((M1, (M2,N))))))"  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
135  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
136  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
20801 
diff
changeset
 | 
137  | 
lappend :: "['a llist, 'a llist] => 'a llist" where  | 
| 19736 | 138  | 
"lappend l n = llist_corec (l,n)  | 
| 
5977
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
139  | 
(split(llist_case (llist_case None (%n1 n2. Some((n1, (LNil,n2)))))  | 
| 
 
9f0c8869cf71
tidied up list definitions, using type 'a option instead of
 
paulson 
parents: 
3842 
diff
changeset
 | 
140  | 
(%l1 l2 n. Some((l1, (l2,n))))))"  | 
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
141  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
142  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
143  | 
text{*Append generates its result by applying f, where
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
144  | 
f((NIL,NIL)) = None  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
145  | 
f((NIL, CONS N1 N2)) = Some((N1, (NIL,N2))  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
146  | 
f((CONS M1 M2, N)) = Some((M1, (M2,N))  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
147  | 
*}  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
148  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
149  | 
text{*
 | 
| 13107 | 150  | 
SHOULD @{text LListD_Fun_CONS_I}, etc., be equations (for rewriting)?
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
151  | 
*}  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
152  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
153  | 
lemmas UN1_I = UNIV_I [THEN UN_I, standard]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
154  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
155  | 
subsubsection{* Simplification *}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
156  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
157  | 
declare option.split [split]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
158  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
159  | 
text{*This justifies using llist in other recursive type definitions*}
 | 
| 23746 | 160  | 
lemma llist_mono:  | 
161  | 
assumes subset: "A \<subseteq> B"  | 
|
162  | 
shows "llist A \<subseteq> llist B"  | 
|
163  | 
proof  | 
|
164  | 
fix x  | 
|
165  | 
assume "x \<in> llist A"  | 
|
166  | 
then show "x \<in> llist B"  | 
|
167  | 
proof coinduct  | 
|
168  | 
case llist  | 
|
169  | 
then show ?case using subset  | 
|
170  | 
by cases blast+  | 
|
171  | 
qed  | 
|
172  | 
qed  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
173  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
174  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
175  | 
lemma llist_unfold: "llist(A) = usum {Numb(0)} (uprod A (llist A))"
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
176  | 
by (fast intro!: llist.intros [unfolded NIL_def CONS_def]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
177  | 
elim: llist.cases [unfolded NIL_def CONS_def])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
178  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
179  | 
|
| 13107 | 180  | 
subsection{* Type checking by coinduction *}
 | 
181  | 
||
182  | 
text {*
 | 
|
183  | 
  {\dots} using @{text list_Fun} THE COINDUCTIVE DEFINITION PACKAGE
 | 
|
184  | 
COULD DO THIS!  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
185  | 
*}  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
186  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
187  | 
lemma llist_coinduct:  | 
| 17841 | 188  | 
"[| M \<in> X; X \<subseteq> list_Fun A (X Un llist(A)) |] ==> M \<in> llist(A)"  | 
189  | 
apply (simp add: list_Fun_def)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
190  | 
apply (erule llist.coinduct)  | 
| 17841 | 191  | 
apply (blast intro: elim:);  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
192  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
193  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
194  | 
lemma list_Fun_NIL_I [iff]: "NIL \<in> list_Fun A X"  | 
| 17841 | 195  | 
by (simp add: list_Fun_def NIL_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
196  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
197  | 
lemma list_Fun_CONS_I [intro!,simp]:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
198  | 
"[| M \<in> A; N \<in> X |] ==> CONS M N \<in> list_Fun A X"  | 
| 17841 | 199  | 
by (simp add: list_Fun_def CONS_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
200  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
201  | 
|
| 13107 | 202  | 
text{*Utilise the ``strong'' part, i.e. @{text "gfp(f)"}*}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
203  | 
lemma list_Fun_llist_I: "M \<in> llist(A) ==> M \<in> list_Fun A (X Un llist(A))"  | 
| 23746 | 204  | 
apply (unfold list_Fun_def)  | 
205  | 
apply (erule llist.cases)  | 
|
206  | 
apply auto  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
207  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
208  | 
|
| 13107 | 209  | 
subsection{* @{text LList_corec} satisfies the desired recurion equation *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
210  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
211  | 
text{*A continuity result?*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
212  | 
lemma CONS_UN1: "CONS M (\<Union>x. f(x)) = (\<Union>x. CONS M (f x))"  | 
| 17841 | 213  | 
apply (simp add: CONS_def In1_UN1 Scons_UN1_y)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
214  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
215  | 
|
| 17841 | 216  | 
lemma CONS_mono: "[| M\<subseteq>M'; N\<subseteq>N' |] ==> CONS M N \<subseteq> CONS M' N'"  | 
217  | 
apply (simp add: CONS_def In1_mono Scons_mono)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
218  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
219  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
220  | 
declare LList_corec_fun_def [THEN def_nat_rec_0, simp]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
221  | 
LList_corec_fun_def [THEN def_nat_rec_Suc, simp]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
222  | 
|
| 19736 | 223  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
224  | 
subsubsection{* The directions of the equality are proved separately *}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
225  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
226  | 
lemma LList_corec_subset1:  | 
| 17841 | 227  | 
"LList_corec a f \<subseteq>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
228  | 
(case f a of None => NIL | Some(z,w) => CONS z (LList_corec w f))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
229  | 
apply (unfold LList_corec_def)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
230  | 
apply (rule UN_least)  | 
| 17841 | 231  | 
apply (case_tac k)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
232  | 
apply (simp_all (no_asm_simp))  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
233  | 
apply (rule allI impI subset_refl [THEN CONS_mono] UNIV_I [THEN UN_upper])+  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
234  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
235  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
236  | 
lemma LList_corec_subset2:  | 
| 17841 | 237  | 
"(case f a of None => NIL | Some(z,w) => CONS z (LList_corec w f)) \<subseteq>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
238  | 
LList_corec a f"  | 
| 17841 | 239  | 
apply (simp add: LList_corec_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
240  | 
apply (simp add: CONS_UN1, safe)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
241  | 
apply (rule_tac a="Suc(?k)" in UN_I, simp, simp)+  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
242  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
243  | 
|
| 13107 | 244  | 
text{*the recursion equation for @{text LList_corec} -- NOT SUITABLE FOR REWRITING!*}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
245  | 
lemma LList_corec:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
246  | 
"LList_corec a f =  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
247  | 
(case f a of None => NIL | Some(z,w) => CONS z (LList_corec w f))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
248  | 
by (rule equalityI LList_corec_subset1 LList_corec_subset2)+  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
249  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
250  | 
text{*definitional version of same*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
251  | 
lemma def_LList_corec:  | 
| 19736 | 252  | 
"[| !!x. h(x) = LList_corec x f |]  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
253  | 
==> h(a) = (case f a of None => NIL | Some(z,w) => CONS z (h w))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
254  | 
by (simp add: LList_corec)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
255  | 
|
| 13107 | 256  | 
text{*A typical use of co-induction to show membership in the @{text gfp}. 
 | 
257  | 
  Bisimulation is @{text "range(%x. LList_corec x f)"} *}
 | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
258  | 
lemma LList_corec_type: "LList_corec a f \<in> llist UNIV"  | 
| 17841 | 259  | 
apply (rule_tac X = "range (%x. LList_corec x ?g)" in llist_coinduct)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
260  | 
apply (rule rangeI, safe)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
261  | 
apply (subst LList_corec, simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
262  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
263  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
264  | 
|
| 13107 | 265  | 
subsection{* @{text llist} equality as a @{text gfp}; the bisimulation principle *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
266  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
267  | 
text{*This theorem is actually used, unlike the many similar ones in ZF*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
268  | 
lemma LListD_unfold: "LListD r = dsum (diag {Numb 0}) (dprod r (LListD r))"
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
269  | 
by (fast intro!: LListD.intros [unfolded NIL_def CONS_def]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
270  | 
elim: LListD.cases [unfolded NIL_def CONS_def])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
271  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
272  | 
lemma LListD_implies_ntrunc_equality [rule_format]:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
273  | 
"\<forall>M N. (M,N) \<in> LListD(diag A) --> ntrunc k M = ntrunc k N"  | 
| 17841 | 274  | 
apply (induct_tac "k" rule: nat_less_induct)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
275  | 
apply (safe del: equalityI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
276  | 
apply (erule LListD.cases)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
277  | 
apply (safe del: equalityI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
278  | 
apply (case_tac "n", simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
279  | 
apply (rename_tac "n'")  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
280  | 
apply (case_tac "n'")  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
281  | 
apply (simp_all add: CONS_def less_Suc_eq)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
282  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
283  | 
|
| 13107 | 284  | 
text{*The domain of the @{text LListD} relation*}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
285  | 
lemma Domain_LListD:  | 
| 17841 | 286  | 
"Domain (LListD(diag A)) \<subseteq> llist(A)"  | 
| 23746 | 287  | 
apply (rule subsetI)  | 
288  | 
apply (erule llist.coinduct)  | 
|
289  | 
apply (simp add: NIL_def CONS_def)  | 
|
290  | 
apply (drule_tac P = "%x. xa \<in> Domain x" in LListD_unfold [THEN subst], auto)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
291  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
292  | 
|
| 13107 | 293  | 
text{*This inclusion justifies the use of coinduction to show @{text "M = N"}*}
 | 
| 17841 | 294  | 
lemma LListD_subset_diag: "LListD(diag A) \<subseteq> diag(llist(A))"  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
295  | 
apply (rule subsetI)  | 
| 17841 | 296  | 
apply (rule_tac p = x in PairE, safe)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
297  | 
apply (rule diag_eqI)  | 
| 17841 | 298  | 
apply (rule LListD_implies_ntrunc_equality [THEN ntrunc_equality], assumption)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
299  | 
apply (erule DomainI [THEN Domain_LListD [THEN subsetD]])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
300  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
301  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
302  | 
|
| 13107 | 303  | 
subsubsection{* Coinduction, using @{text LListD_Fun} *}
 | 
304  | 
||
305  | 
text {* THE COINDUCTIVE DEFINITION PACKAGE COULD DO THIS! *}
 | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
306  | 
|
| 17841 | 307  | 
lemma LListD_Fun_mono: "A\<subseteq>B ==> LListD_Fun r A \<subseteq> LListD_Fun r B"  | 
308  | 
apply (simp add: LListD_Fun_def)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
309  | 
apply (assumption | rule basic_monos)+  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
310  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
311  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
312  | 
lemma LListD_coinduct:  | 
| 17841 | 313  | 
"[| M \<in> X; X \<subseteq> LListD_Fun r (X Un LListD(r)) |] ==> M \<in> LListD(r)"  | 
| 23746 | 314  | 
apply (cases M)  | 
| 17841 | 315  | 
apply (simp add: LListD_Fun_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
316  | 
apply (erule LListD.coinduct)  | 
| 17841 | 317  | 
apply (auto );  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
318  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
319  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
320  | 
lemma LListD_Fun_NIL_I: "(NIL,NIL) \<in> LListD_Fun r s"  | 
| 17841 | 321  | 
by (simp add: LListD_Fun_def NIL_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
322  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
323  | 
lemma LListD_Fun_CONS_I:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
324  | 
"[| x\<in>A; (M,N):s |] ==> (CONS x M, CONS x N) \<in> LListD_Fun (diag A) s"  | 
| 17841 | 325  | 
by (simp add: LListD_Fun_def CONS_def, blast)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
326  | 
|
| 13107 | 327  | 
text{*Utilise the "strong" part, i.e. @{text "gfp(f)"}*}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
328  | 
lemma LListD_Fun_LListD_I:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
329  | 
"M \<in> LListD(r) ==> M \<in> LListD_Fun r (X Un LListD(r))"  | 
| 23746 | 330  | 
apply (cases M)  | 
331  | 
apply (simp add: LListD_Fun_def)  | 
|
332  | 
apply (erule LListD.cases)  | 
|
333  | 
apply auto  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
334  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
335  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
336  | 
|
| 13107 | 337  | 
text{*This converse inclusion helps to strengthen @{text LList_equalityI}*}
 | 
| 17841 | 338  | 
lemma diag_subset_LListD: "diag(llist(A)) \<subseteq> LListD(diag A)"  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
339  | 
apply (rule subsetI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
340  | 
apply (erule LListD_coinduct)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
341  | 
apply (rule subsetI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
342  | 
apply (erule diagE)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
343  | 
apply (erule ssubst)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
344  | 
apply (erule llist.cases)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
345  | 
apply (simp_all add: diagI LListD_Fun_NIL_I LListD_Fun_CONS_I)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
346  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
347  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
348  | 
lemma LListD_eq_diag: "LListD(diag A) = diag(llist(A))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
349  | 
apply (rule equalityI LListD_subset_diag diag_subset_LListD)+  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
350  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
351  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
352  | 
lemma LListD_Fun_diag_I: "M \<in> llist(A) ==> (M,M) \<in> LListD_Fun (diag A) (X Un diag(llist(A)))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
353  | 
apply (rule LListD_eq_diag [THEN subst])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
354  | 
apply (rule LListD_Fun_LListD_I)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
355  | 
apply (simp add: LListD_eq_diag diagI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
356  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
357  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
358  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
359  | 
subsubsection{* To show two LLists are equal, exhibit a bisimulation! 
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
360  | 
[also admits true equality]  | 
| 13107 | 361  | 
   Replace @{text A} by some particular set, like @{text "{x. True}"}??? *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
362  | 
lemma LList_equalityI:  | 
| 17841 | 363  | 
"[| (M,N) \<in> r; r \<subseteq> LListD_Fun (diag A) (r Un diag(llist(A))) |]  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
364  | 
==> M=N"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
365  | 
apply (rule LListD_subset_diag [THEN subsetD, THEN diagE])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
366  | 
apply (erule LListD_coinduct)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
367  | 
apply (simp add: LListD_eq_diag, safe)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
368  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
369  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
370  | 
|
| 13107 | 371  | 
subsection{* Finality of @{text "llist(A)"}: Uniqueness of functions defined by corecursion *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
372  | 
|
| 13107 | 373  | 
text{*We must remove @{text Pair_eq} because it may turn an instance of reflexivity
 | 
374  | 
  @{text "(h1 b, h2 b) = (h1 ?x17, h2 ?x17)"} into a conjunction! 
 | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
375  | 
(or strengthen the Solver?)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
376  | 
*}  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
377  | 
declare Pair_eq [simp del]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
378  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
379  | 
text{*abstract proof using a bisimulation*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
380  | 
lemma LList_corec_unique:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
381  | 
"[| !!x. h1(x) = (case f x of None => NIL | Some(z,w) => CONS z (h1 w));  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
382  | 
!!x. h2(x) = (case f x of None => NIL | Some(z,w) => CONS z (h2 w)) |]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
383  | 
==> h1=h2"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
384  | 
apply (rule ext)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
385  | 
txt{*next step avoids an unknown (and flexflex pair) in simplification*}
 | 
| 17841 | 386  | 
apply (rule_tac A = UNIV and r = "range(%u. (h1 u, h2 u))"  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
387  | 
in LList_equalityI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
388  | 
apply (rule rangeI, safe)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
389  | 
apply (simp add: LListD_Fun_NIL_I UNIV_I [THEN LListD_Fun_CONS_I])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
390  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
391  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
392  | 
lemma equals_LList_corec:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
393  | 
"[| !!x. h(x) = (case f x of None => NIL | Some(z,w) => CONS z (h w)) |]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
394  | 
==> h = (%x. LList_corec x f)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
395  | 
by (simp add: LList_corec_unique LList_corec)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
396  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
397  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
398  | 
subsubsection{*Obsolete proof of @{text LList_corec_unique}: 
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
399  | 
complete induction, not coinduction *}  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
400  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
401  | 
lemma ntrunc_one_CONS [simp]: "ntrunc (Suc 0) (CONS M N) = {}"
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
402  | 
by (simp add: CONS_def ntrunc_one_In1)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
403  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
404  | 
lemma ntrunc_CONS [simp]:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
405  | 
"ntrunc (Suc(Suc(k))) (CONS M N) = CONS (ntrunc k M) (ntrunc k N)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
406  | 
by (simp add: CONS_def)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
407  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
408  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
409  | 
lemma  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
410  | 
assumes prem1:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
411  | 
"!!x. h1 x = (case f x of None => NIL | Some(z,w) => CONS z (h1 w))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
412  | 
and prem2:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
413  | 
"!!x. h2 x = (case f x of None => NIL | Some(z,w) => CONS z (h2 w))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
414  | 
shows "h1=h2"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
415  | 
apply (rule ntrunc_equality [THEN ext])  | 
| 17841 | 416  | 
apply (rule_tac x = x in spec)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
417  | 
apply (induct_tac "k" rule: nat_less_induct)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
418  | 
apply (rename_tac "n")  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
419  | 
apply (rule allI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
420  | 
apply (subst prem1)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
421  | 
apply (subst prem2, simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
422  | 
apply (intro strip)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
423  | 
apply (case_tac "n")  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
424  | 
apply (rename_tac [2] "m")  | 
| 17841 | 425  | 
apply (case_tac [2] "m", simp_all)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
426  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
427  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
428  | 
|
| 13107 | 429  | 
subsection{*Lconst: defined directly by @{text lfp} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
430  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
431  | 
text{*But it could be defined by corecursion.*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
432  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
433  | 
lemma Lconst_fun_mono: "mono(CONS(M))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
434  | 
apply (rule monoI subset_refl CONS_mono)+  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
435  | 
apply assumption  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
436  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
437  | 
|
| 13107 | 438  | 
text{* @{text "Lconst(M) = CONS M (Lconst M)"} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
439  | 
lemmas Lconst = Lconst_fun_mono [THEN Lconst_def [THEN def_lfp_unfold]]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
440  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
441  | 
text{*A typical use of co-induction to show membership in the gfp.
 | 
| 13107 | 442  | 
  The containing set is simply the singleton @{text "{Lconst(M)}"}. *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
443  | 
lemma Lconst_type: "M\<in>A ==> Lconst(M): llist(A)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
444  | 
apply (rule singletonI [THEN llist_coinduct], safe)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
445  | 
apply (rule_tac P = "%u. u \<in> ?A" in Lconst [THEN ssubst])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
446  | 
apply (assumption | rule list_Fun_CONS_I singletonI UnI1)+  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
447  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
448  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
449  | 
lemma Lconst_eq_LList_corec: "Lconst(M) = LList_corec M (%x. Some(x,x))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
450  | 
apply (rule equals_LList_corec [THEN fun_cong], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
451  | 
apply (rule Lconst)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
452  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
453  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
454  | 
text{*Thus we could have used gfp in the definition of Lconst*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
455  | 
lemma gfp_Lconst_eq_LList_corec: "gfp(%N. CONS M N) = LList_corec M (%x. Some(x,x))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
456  | 
apply (rule equals_LList_corec [THEN fun_cong], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
457  | 
apply (rule Lconst_fun_mono [THEN gfp_unfold])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
458  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
459  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
460  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
461  | 
subsection{* Isomorphisms *}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
462  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
463  | 
lemma LListI: "x \<in> llist (range Leaf) ==> x \<in> LList"  | 
| 17841 | 464  | 
by (simp add: LList_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
465  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
466  | 
lemma LListD: "x \<in> LList ==> x \<in> llist (range Leaf)"  | 
| 17841 | 467  | 
by (simp add: LList_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
468  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
469  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
470  | 
subsubsection{* Distinctness of constructors *}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
471  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
472  | 
lemma LCons_not_LNil [iff]: "~ LCons x xs = LNil"  | 
| 17841 | 473  | 
apply (simp add: LNil_def LCons_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
474  | 
apply (subst Abs_LList_inject)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
475  | 
apply (rule llist.intros CONS_not_NIL rangeI LListI Rep_LList [THEN LListD])+  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
476  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
477  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
478  | 
lemmas LNil_not_LCons [iff] = LCons_not_LNil [THEN not_sym, standard]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
479  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
480  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
481  | 
subsubsection{* llist constructors *}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
482  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
483  | 
lemma Rep_LList_LNil: "Rep_LList LNil = NIL"  | 
| 17841 | 484  | 
apply (simp add: LNil_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
485  | 
apply (rule llist.NIL_I [THEN LListI, THEN Abs_LList_inverse])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
486  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
487  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
488  | 
lemma Rep_LList_LCons: "Rep_LList(LCons x l) = CONS (Leaf x) (Rep_LList l)"  | 
| 17841 | 489  | 
apply (simp add: LCons_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
490  | 
apply (rule llist.CONS_I [THEN LListI, THEN Abs_LList_inverse]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
491  | 
rangeI Rep_LList [THEN LListD])+  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
492  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
493  | 
|
| 13107 | 494  | 
subsubsection{* Injectiveness of @{text CONS} and @{text LCons} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
495  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
496  | 
lemma CONS_CONS_eq2: "(CONS M N=CONS M' N') = (M=M' & N=N')"  | 
| 17841 | 497  | 
apply (simp add: CONS_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
498  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
499  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
500  | 
lemmas CONS_inject = CONS_CONS_eq [THEN iffD1, THEN conjE, standard]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
501  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
502  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
503  | 
text{*For reasoning about abstract llist constructors*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
504  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
505  | 
declare Rep_LList [THEN LListD, intro] LListI [intro]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
506  | 
declare llist.intros [intro]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
507  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
508  | 
lemma LCons_LCons_eq [iff]: "(LCons x xs=LCons y ys) = (x=y & xs=ys)"  | 
| 17841 | 509  | 
apply (simp add: LCons_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
510  | 
apply (subst Abs_LList_inject)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
511  | 
apply (auto simp add: Rep_LList_inject)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
512  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
513  | 
|
| 13524 | 514  | 
lemma CONS_D2: "CONS M N \<in> llist(A) ==> M \<in> A & N \<in> llist(A)"  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
515  | 
apply (erule llist.cases)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
516  | 
apply (erule CONS_neq_NIL, fast)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
517  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
518  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
519  | 
|
| 13107 | 520  | 
subsection{* Reasoning about @{text "llist(A)"} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
521  | 
|
| 13107 | 522  | 
text{*A special case of @{text list_equality} for functions over lazy lists*}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
523  | 
lemma LList_fun_equalityI:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
524  | 
"[| M \<in> llist(A); g(NIL): llist(A);  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
525  | 
f(NIL)=g(NIL);  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
526  | 
!!x l. [| x\<in>A; l \<in> llist(A) |] ==>  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
527  | 
(f(CONS x l),g(CONS x l)) \<in>  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
528  | 
LListD_Fun (diag A) ((%u.(f(u),g(u)))`llist(A) Un  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
529  | 
diag(llist(A)))  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
530  | 
|] ==> f(M) = g(M)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
531  | 
apply (rule LList_equalityI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
532  | 
apply (erule imageI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
533  | 
apply (rule image_subsetI)  | 
| 23746 | 534  | 
apply (erule_tac a=x in llist.cases)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
535  | 
apply (erule ssubst, erule ssubst, erule LListD_Fun_diag_I, blast)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
536  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
537  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
538  | 
|
| 13107 | 539  | 
subsection{* The functional @{text Lmap} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
540  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
541  | 
lemma Lmap_NIL [simp]: "Lmap f NIL = NIL"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
542  | 
by (rule Lmap_def [THEN def_LList_corec, THEN trans], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
543  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
544  | 
lemma Lmap_CONS [simp]: "Lmap f (CONS M N) = CONS (f M) (Lmap f N)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
545  | 
by (rule Lmap_def [THEN def_LList_corec, THEN trans], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
546  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
547  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
548  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
549  | 
text{*Another type-checking proof by coinduction*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
550  | 
lemma Lmap_type:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
551  | 
"[| M \<in> llist(A); !!x. x\<in>A ==> f(x):B |] ==> Lmap f M \<in> llist(B)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
552  | 
apply (erule imageI [THEN llist_coinduct], safe)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
553  | 
apply (erule llist.cases, simp_all)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
554  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
555  | 
|
| 13107 | 556  | 
text{*This type checking rule synthesises a sufficiently large set for @{text f}*}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
557  | 
lemma Lmap_type2: "M \<in> llist(A) ==> Lmap f M \<in> llist(f`A)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
558  | 
apply (erule Lmap_type)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
559  | 
apply (erule imageI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
560  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
561  | 
|
| 13107 | 562  | 
subsubsection{* Two easy results about @{text Lmap} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
563  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
564  | 
lemma Lmap_compose: "M \<in> llist(A) ==> Lmap (f o g) M = Lmap f (Lmap g M)"  | 
| 17841 | 565  | 
apply (simp add: o_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
566  | 
apply (drule imageI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
567  | 
apply (erule LList_equalityI, safe)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
568  | 
apply (erule llist.cases, simp_all)  | 
| 
26793
 
e36a92ff543e
Instantiated some rules to avoid problems with HO unification.
 
berghofe 
parents: 
23746 
diff
changeset
 | 
569  | 
apply (rule LListD_Fun_NIL_I imageI UnI1 rangeI [THEN LListD_Fun_CONS_I, of _ _ _ f])+  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
570  | 
apply assumption  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
571  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
572  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
573  | 
lemma Lmap_ident: "M \<in> llist(A) ==> Lmap (%x. x) M = M"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
574  | 
apply (drule imageI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
575  | 
apply (erule LList_equalityI, safe)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
576  | 
apply (erule llist.cases, simp_all)  | 
| 
26793
 
e36a92ff543e
Instantiated some rules to avoid problems with HO unification.
 
berghofe 
parents: 
23746 
diff
changeset
 | 
577  | 
apply (rule LListD_Fun_NIL_I imageI UnI1 rangeI [THEN LListD_Fun_CONS_I, of _ _ _ "%x. x"])+  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
578  | 
apply assumption  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
579  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
580  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
581  | 
|
| 13107 | 582  | 
subsection{* @{text Lappend} -- its two arguments cause some complications! *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
583  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
584  | 
lemma Lappend_NIL_NIL [simp]: "Lappend NIL NIL = NIL"  | 
| 17841 | 585  | 
apply (simp add: Lappend_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
586  | 
apply (rule LList_corec [THEN trans], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
587  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
588  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
589  | 
lemma Lappend_NIL_CONS [simp]:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
590  | 
"Lappend NIL (CONS N N') = CONS N (Lappend NIL N')"  | 
| 17841 | 591  | 
apply (simp add: Lappend_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
592  | 
apply (rule LList_corec [THEN trans], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
593  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
594  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
595  | 
lemma Lappend_CONS [simp]:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
596  | 
"Lappend (CONS M M') N = CONS M (Lappend M' N)"  | 
| 17841 | 597  | 
apply (simp add: Lappend_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
598  | 
apply (rule LList_corec [THEN trans], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
599  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
600  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
601  | 
declare llist.intros [simp] LListD_Fun_CONS_I [simp]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
602  | 
range_eqI [simp] image_eqI [simp]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
603  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
604  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
605  | 
lemma Lappend_NIL [simp]: "M \<in> llist(A) ==> Lappend NIL M = M"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
606  | 
by (erule LList_fun_equalityI, simp_all)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
607  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
608  | 
lemma Lappend_NIL2: "M \<in> llist(A) ==> Lappend M NIL = M"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
609  | 
by (erule LList_fun_equalityI, simp_all)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
610  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
611  | 
|
| 13107 | 612  | 
subsubsection{* Alternative type-checking proofs for @{text Lappend} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
613  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
614  | 
text{*weak co-induction: bisimulation and case analysis on both variables*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
615  | 
lemma Lappend_type: "[| M \<in> llist(A); N \<in> llist(A) |] ==> Lappend M N \<in> llist(A)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
616  | 
apply (rule_tac X = "\<Union>u\<in>llist (A) . \<Union>v \<in> llist (A) . {Lappend u v}" in llist_coinduct)
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
617  | 
apply fast  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
618  | 
apply safe  | 
| 23746 | 619  | 
apply (erule_tac a = u in llist.cases)  | 
620  | 
apply (erule_tac a = v in llist.cases, simp_all, blast)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
621  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
622  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
623  | 
text{*strong co-induction: bisimulation and case analysis on one variable*}
 | 
| 13524 | 624  | 
lemma Lappend_type': "[| M \<in> llist(A); N \<in> llist(A) |] ==> Lappend M N \<in> llist(A)"  | 
| 17841 | 625  | 
apply (rule_tac X = "(%u. Lappend u N) `llist (A)" in llist_coinduct)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
626  | 
apply (erule imageI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
627  | 
apply (rule image_subsetI)  | 
| 23746 | 628  | 
apply (erule_tac a = x in llist.cases)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
629  | 
apply (simp add: list_Fun_llist_I, simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
630  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
631  | 
|
| 13107 | 632  | 
subsection{* Lazy lists as the type @{text "'a llist"} -- strongly typed versions of above *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
633  | 
|
| 13107 | 634  | 
subsubsection{* @{text llist_case}: case analysis for @{text "'a llist"} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
635  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
636  | 
declare LListI [THEN Abs_LList_inverse, simp]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
637  | 
declare Rep_LList_inverse [simp]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
638  | 
declare Rep_LList [THEN LListD, simp]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
639  | 
declare rangeI [simp] inj_Leaf [simp]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
640  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
641  | 
lemma llist_case_LNil [simp]: "llist_case c d LNil = c"  | 
| 17841 | 642  | 
by (simp add: llist_case_def LNil_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
643  | 
|
| 17841 | 644  | 
lemma llist_case_LCons [simp]: "llist_case c d (LCons M N) = d M N"  | 
645  | 
by (simp add: llist_case_def LCons_def)  | 
|
646  | 
||
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
647  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
648  | 
text{*Elimination is case analysis, not induction.*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
649  | 
lemma llistE: "[| l=LNil ==> P; !!x l'. l=LCons x l' ==> P |] ==> P"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
650  | 
apply (rule Rep_LList [THEN LListD, THEN llist.cases])  | 
| 17841 | 651  | 
apply (simp add: Rep_LList_LNil [symmetric] Rep_LList_inject, blast)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
652  | 
apply (erule LListI [THEN Rep_LList_cases], clarify)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
653  | 
apply (simp add: Rep_LList_LCons [symmetric] Rep_LList_inject, blast)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
654  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
655  | 
|
| 13107 | 656  | 
subsubsection{* @{text llist_corec}: corecursion for @{text "'a llist"} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
657  | 
|
| 13107 | 658  | 
text{*Lemma for the proof of @{text llist_corec}*}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
659  | 
lemma LList_corec_type2:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
660  | 
"LList_corec a  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
661  | 
(%z. case f z of None => None | Some(v,w) => Some(Leaf(v),w))  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
662  | 
\<in> llist(range Leaf)"  | 
| 17841 | 663  | 
apply (rule_tac X = "range (%x. LList_corec x ?g)" in llist_coinduct)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
664  | 
apply (rule rangeI, safe)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
665  | 
apply (subst LList_corec, force)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
666  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
667  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
668  | 
lemma llist_corec:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
669  | 
"llist_corec a f =  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
670  | 
(case f a of None => LNil | Some(z,w) => LCons z (llist_corec w f))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
671  | 
apply (unfold llist_corec_def LNil_def LCons_def)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
672  | 
apply (subst LList_corec)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
673  | 
apply (case_tac "f a")  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
674  | 
apply (simp add: LList_corec_type2)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
675  | 
apply (force simp add: LList_corec_type2)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
676  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
677  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
678  | 
text{*definitional version of same*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
679  | 
lemma def_llist_corec:  | 
| 19736 | 680  | 
"[| !!x. h(x) = llist_corec x f |] ==>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
681  | 
h(a) = (case f a of None => LNil | Some(z,w) => LCons z (h w))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
682  | 
by (simp add: llist_corec)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
683  | 
|
| 13107 | 684  | 
subsection{* Proofs about type @{text "'a llist"} functions *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
685  | 
|
| 13107 | 686  | 
subsection{* Deriving @{text llist_equalityI} -- @{text llist} equality is a bisimulation *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
687  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
688  | 
lemma LListD_Fun_subset_Times_llist:  | 
| 17841 | 689  | 
"r \<subseteq> (llist A) <*> (llist A)  | 
690  | 
==> LListD_Fun (diag A) r \<subseteq> (llist A) <*> (llist A)"  | 
|
| 15481 | 691  | 
by (auto simp add: LListD_Fun_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
692  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
693  | 
lemma subset_Times_llist:  | 
| 17841 | 694  | 
"prod_fun Rep_LList Rep_LList ` r \<subseteq>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
695  | 
(llist(range Leaf)) <*> (llist(range Leaf))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
696  | 
by (blast intro: Rep_LList [THEN LListD])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
697  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
698  | 
lemma prod_fun_lemma:  | 
| 17841 | 699  | 
"r \<subseteq> (llist(range Leaf)) <*> (llist(range Leaf))  | 
700  | 
==> prod_fun (Rep_LList o Abs_LList) (Rep_LList o Abs_LList) ` r \<subseteq> r"  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
701  | 
apply safe  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
702  | 
apply (erule subsetD [THEN SigmaE2], assumption)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
703  | 
apply (simp add: LListI [THEN Abs_LList_inverse])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
704  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
705  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
706  | 
lemma prod_fun_range_eq_diag:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
707  | 
"prod_fun Rep_LList Rep_LList ` range(%x. (x, x)) =  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
708  | 
diag(llist(range Leaf))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
709  | 
apply (rule equalityI, blast)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
710  | 
apply (fast elim: LListI [THEN Abs_LList_inverse, THEN subst])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
711  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
712  | 
|
| 13107 | 713  | 
text{*Used with @{text lfilter}*}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
714  | 
lemma llistD_Fun_mono:  | 
| 17841 | 715  | 
"A\<subseteq>B ==> llistD_Fun A \<subseteq> llistD_Fun B"  | 
716  | 
apply (simp add: llistD_Fun_def prod_fun_def, auto)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
717  | 
apply (rule image_eqI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
718  | 
prefer 2 apply (blast intro: rev_subsetD [OF _ LListD_Fun_mono], force)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
719  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
720  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
721  | 
subsubsection{* To show two llists are equal, exhibit a bisimulation! 
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
722  | 
[also admits true equality] *}  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
723  | 
lemma llist_equalityI:  | 
| 17841 | 724  | 
"[| (l1,l2) \<in> r; r \<subseteq> llistD_Fun(r Un range(%x.(x,x))) |] ==> l1=l2"  | 
725  | 
apply (simp add: llistD_Fun_def)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
726  | 
apply (rule Rep_LList_inject [THEN iffD1])  | 
| 17841 | 727  | 
apply (rule_tac r = "prod_fun Rep_LList Rep_LList `r" and A = "range (Leaf)" in LList_equalityI)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
728  | 
apply (erule prod_fun_imageI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
729  | 
apply (erule image_mono [THEN subset_trans])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
730  | 
apply (rule image_compose [THEN subst])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
731  | 
apply (rule prod_fun_compose [THEN subst])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
732  | 
apply (subst image_Un)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
733  | 
apply (subst prod_fun_range_eq_diag)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
734  | 
apply (rule LListD_Fun_subset_Times_llist [THEN prod_fun_lemma])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
735  | 
apply (rule subset_Times_llist [THEN Un_least])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
736  | 
apply (rule diag_subset_Times)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
737  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
738  | 
|
| 13107 | 739  | 
subsubsection{* Rules to prove the 2nd premise of @{text llist_equalityI} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
740  | 
lemma llistD_Fun_LNil_I [simp]: "(LNil,LNil) \<in> llistD_Fun(r)"  | 
| 17841 | 741  | 
apply (simp add: llistD_Fun_def LNil_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
742  | 
apply (rule LListD_Fun_NIL_I [THEN prod_fun_imageI])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
743  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
744  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
745  | 
lemma llistD_Fun_LCons_I [simp]:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
746  | 
"(l1,l2):r ==> (LCons x l1, LCons x l2) \<in> llistD_Fun(r)"  | 
| 17841 | 747  | 
apply (simp add: llistD_Fun_def LCons_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
748  | 
apply (rule rangeI [THEN LListD_Fun_CONS_I, THEN prod_fun_imageI])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
749  | 
apply (erule prod_fun_imageI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
750  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
751  | 
|
| 13107 | 752  | 
text{*Utilise the "strong" part, i.e. @{text "gfp(f)"}*}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
753  | 
lemma llistD_Fun_range_I: "(l,l) \<in> llistD_Fun(r Un range(%x.(x,x)))"  | 
| 17841 | 754  | 
apply (simp add: llistD_Fun_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
755  | 
apply (rule Rep_LList_inverse [THEN subst])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
756  | 
apply (rule prod_fun_imageI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
757  | 
apply (subst image_Un)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
758  | 
apply (subst prod_fun_range_eq_diag)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
759  | 
apply (rule Rep_LList [THEN LListD, THEN LListD_Fun_diag_I])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
760  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
761  | 
|
| 13107 | 762  | 
text{*A special case of @{text list_equality} for functions over lazy lists*}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
763  | 
lemma llist_fun_equalityI:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
764  | 
"[| f(LNil)=g(LNil);  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
765  | 
!!x l. (f(LCons x l),g(LCons x l))  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
766  | 
\<in> llistD_Fun(range(%u. (f(u),g(u))) Un range(%v. (v,v)))  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
767  | 
|] ==> f(l) = (g(l :: 'a llist) :: 'b llist)"  | 
| 17841 | 768  | 
apply (rule_tac r = "range (%u. (f (u),g (u)))" in llist_equalityI)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
769  | 
apply (rule rangeI, clarify)  | 
| 17841 | 770  | 
apply (rule_tac l = u in llistE)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
771  | 
apply (simp_all add: llistD_Fun_range_I)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
772  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
773  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
774  | 
|
| 13107 | 775  | 
subsection{* The functional @{text lmap} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
776  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
777  | 
lemma lmap_LNil [simp]: "lmap f LNil = LNil"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
778  | 
by (rule lmap_def [THEN def_llist_corec, THEN trans], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
779  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
780  | 
lemma lmap_LCons [simp]: "lmap f (LCons M N) = LCons (f M) (lmap f N)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
781  | 
by (rule lmap_def [THEN def_llist_corec, THEN trans], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
782  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
783  | 
|
| 13107 | 784  | 
subsubsection{* Two easy results about @{text lmap} *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
785  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
786  | 
lemma lmap_compose [simp]: "lmap (f o g) l = lmap f (lmap g l)"  | 
| 17841 | 787  | 
by (rule_tac l = l in llist_fun_equalityI, simp_all)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
788  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
789  | 
lemma lmap_ident [simp]: "lmap (%x. x) l = l"  | 
| 17841 | 790  | 
by (rule_tac l = l in llist_fun_equalityI, simp_all)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
791  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
792  | 
|
| 13107 | 793  | 
subsection{* iterates -- @{text llist_fun_equalityI} cannot be used! *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
794  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
795  | 
lemma iterates: "iterates f x = LCons x (iterates f (f x))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
796  | 
by (rule iterates_def [THEN def_llist_corec, THEN trans], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
797  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
798  | 
lemma lmap_iterates [simp]: "lmap f (iterates f x) = iterates f (f x)"  | 
| 17841 | 799  | 
apply (rule_tac r = "range (%u. (lmap f (iterates f u),iterates f (f u)))" in llist_equalityI)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
800  | 
apply (rule rangeI, safe)  | 
| 17841 | 801  | 
apply (rule_tac x1 = "f (u)" in iterates [THEN ssubst])  | 
802  | 
apply (rule_tac x1 = u in iterates [THEN ssubst], simp)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
803  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
804  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
805  | 
lemma iterates_lmap: "iterates f x = LCons x (lmap f (iterates f x))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
806  | 
apply (subst lmap_iterates)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
807  | 
apply (rule iterates)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
808  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
809  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
810  | 
subsection{* A rather complex proof about iterates -- cf Andy Pitts *}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
811  | 
|
| 13107 | 812  | 
subsubsection{* Two lemmas about @{text "natrec n x (%m. g)"}, which is essentially
 | 
813  | 
  @{text "(g^n)(x)"} *}
 | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
814  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
815  | 
lemma fun_power_lmap: "nat_rec (LCons b l) (%m. lmap(f)) n =  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
816  | 
LCons (nat_rec b (%m. f) n) (nat_rec l (%m. lmap(f)) n)"  | 
| 17841 | 817  | 
by (induct_tac "n", simp_all)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
818  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
819  | 
lemma fun_power_Suc: "nat_rec (g x) (%m. g) n = nat_rec x (%m. g) (Suc n)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
820  | 
by (induct_tac "n", simp_all)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
821  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
822  | 
lemmas Pair_cong = refl [THEN cong, THEN cong, of concl: Pair]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
823  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
824  | 
|
| 13107 | 825  | 
text{*The bisimulation consists of @{text "{(lmap(f)^n (h(u)), lmap(f)^n (iterates(f,u)))}"}
 | 
826  | 
  for all @{text u} and all @{text "n::nat"}.*}
 | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
827  | 
lemma iterates_equality:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
828  | 
"(!!x. h(x) = LCons x (lmap f (h x))) ==> h = iterates(f)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
829  | 
apply (rule ext)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
830  | 
apply (rule_tac  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
831  | 
r = "\<Union>u. range (%n. (nat_rec (h u) (%m y. lmap f y) n,  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
832  | 
nat_rec (iterates f u) (%m y. lmap f y) n))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
833  | 
in llist_equalityI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
834  | 
apply (rule UN1_I range_eqI Pair_cong nat_rec_0 [symmetric])+  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
835  | 
apply clarify  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
836  | 
apply (subst iterates, atomize)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
837  | 
apply (drule_tac x=u in spec)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
838  | 
apply (erule ssubst)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
839  | 
apply (subst fun_power_lmap)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
840  | 
apply (subst fun_power_lmap)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
841  | 
apply (rule llistD_Fun_LCons_I)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
842  | 
apply (rule lmap_iterates [THEN subst])  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
843  | 
apply (subst fun_power_Suc)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
844  | 
apply (subst fun_power_Suc, blast)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
845  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
846  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
847  | 
|
| 13107 | 848  | 
subsection{* @{text lappend} -- its two arguments cause some complications! *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
849  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
850  | 
lemma lappend_LNil_LNil [simp]: "lappend LNil LNil = LNil"  | 
| 17841 | 851  | 
apply (simp add: lappend_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
852  | 
apply (rule llist_corec [THEN trans], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
853  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
854  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
855  | 
lemma lappend_LNil_LCons [simp]:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
856  | 
"lappend LNil (LCons l l') = LCons l (lappend LNil l')"  | 
| 17841 | 857  | 
apply (simp add: lappend_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
858  | 
apply (rule llist_corec [THEN trans], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
859  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
860  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
861  | 
lemma lappend_LCons [simp]:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
862  | 
"lappend (LCons l l') N = LCons l (lappend l' N)"  | 
| 17841 | 863  | 
apply (simp add: lappend_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
864  | 
apply (rule llist_corec [THEN trans], simp)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
865  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
866  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
867  | 
lemma lappend_LNil [simp]: "lappend LNil l = l"  | 
| 17841 | 868  | 
by (rule_tac l = l in llist_fun_equalityI, simp_all)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
869  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
870  | 
lemma lappend_LNil2 [simp]: "lappend l LNil = l"  | 
| 17841 | 871  | 
by (rule_tac l = l in llist_fun_equalityI, simp_all)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
872  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
873  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
874  | 
text{*The infinite first argument blocks the second*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
875  | 
lemma lappend_iterates [simp]: "lappend (iterates f x) N = iterates f x"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
876  | 
apply (rule_tac r = "range (%u. (lappend (iterates f u) N,iterates f u))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
877  | 
in llist_equalityI)  | 
| 17841 | 878  | 
apply (rule rangeI, safe)  | 
| 15944 | 879  | 
apply (subst (1 2) iterates)  | 
880  | 
apply simp  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
881  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
882  | 
|
| 13107 | 883  | 
subsubsection{* Two proofs that @{text lmap} distributes over lappend *}
 | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
884  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
885  | 
text{*Long proof requiring case analysis on both both arguments*}
 | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
886  | 
lemma lmap_lappend_distrib:  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
887  | 
"lmap f (lappend l n) = lappend (lmap f l) (lmap f n)"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
888  | 
apply (rule_tac r = "\<Union>n. range (%l. (lmap f (lappend l n),  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
889  | 
lappend (lmap f l) (lmap f n)))"  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
890  | 
in llist_equalityI)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
891  | 
apply (rule UN1_I)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
892  | 
apply (rule rangeI, safe)  | 
| 17841 | 893  | 
apply (rule_tac l = l in llistE)  | 
894  | 
apply (rule_tac l = n in llistE, simp_all)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
895  | 
apply (blast intro: llistD_Fun_LCons_I)  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
896  | 
done  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
897  | 
|
| 13107 | 898  | 
text{*Shorter proof of theorem above using @{text llist_equalityI} as strong coinduction*}
 | 
| 13524 | 899  | 
lemma lmap_lappend_distrib':  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
900  | 
"lmap f (lappend l n) = lappend (lmap f l) (lmap f n)"  | 
| 17841 | 901  | 
by (rule_tac l = l in llist_fun_equalityI, auto)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
902  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
903  | 
text{*Without strong coinduction, three case analyses might be needed*}
 | 
| 13524 | 904  | 
lemma lappend_assoc': "lappend (lappend l1 l2) l3 = lappend l1 (lappend l2 l3)"  | 
| 17841 | 905  | 
by (rule_tac l = l1 in llist_fun_equalityI, auto)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10834 
diff
changeset
 | 
906  | 
|
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
907  | 
end  |