| author | wenzelm | 
| Sun, 10 Feb 2019 19:07:53 +0100 | |
| changeset 69798 | f610115ca3d0 | 
| parent 62020 | 5d208fd2507d | 
| permissions | -rw-r--r-- | 
| 17456 | 1  | 
(* Title: CCL/Hered.thy  | 
| 1474 | 2  | 
Author: Martin Coen  | 
| 0 | 3  | 
Copyright 1993 University of Cambridge  | 
4  | 
*)  | 
|
5  | 
||
| 60770 | 6  | 
section \<open>Hereditary Termination -- cf. Martin Lo\"f\<close>  | 
| 17456 | 7  | 
|
8  | 
theory Hered  | 
|
9  | 
imports Type  | 
|
10  | 
begin  | 
|
11  | 
||
| 60770 | 12  | 
text \<open>  | 
| 62020 | 13  | 
Note that this is based on an untyped equality and so \<open>lam  | 
14  | 
x. b(x)\<close> is only hereditarily terminating if \<open>ALL x. b(x)\<close>  | 
|
| 17456 | 15  | 
is. Not so useful for functions!  | 
| 60770 | 16  | 
\<close>  | 
| 0 | 17  | 
|
| 58977 | 18  | 
definition HTTgen :: "i set \<Rightarrow> i set" where  | 
| 42156 | 19  | 
"HTTgen(R) ==  | 
| 58977 | 20  | 
    {t. t=true | t=false | (EX a b. t= <a, b> \<and> a : R \<and> b : R) |
 | 
21  | 
(EX f. t = lam x. f(x) \<and> (ALL x. f(x) : R))}"  | 
|
| 0 | 22  | 
|
| 42156 | 23  | 
definition HTT :: "i set"  | 
24  | 
where "HTT == gfp(HTTgen)"  | 
|
| 17456 | 25  | 
|
| 20140 | 26  | 
|
| 60770 | 27  | 
subsection \<open>Hereditary Termination\<close>  | 
| 20140 | 28  | 
|
| 58977 | 29  | 
lemma HTTgen_mono: "mono(\<lambda>X. HTTgen(X))"  | 
| 20140 | 30  | 
apply (unfold HTTgen_def)  | 
31  | 
apply (rule monoI)  | 
|
32  | 
apply blast  | 
|
33  | 
done  | 
|
34  | 
||
35  | 
lemma HTTgenXH:  | 
|
| 58977 | 36  | 
"t : HTTgen(A) \<longleftrightarrow> t=true | t=false | (EX a b. t=<a,b> \<and> a : A \<and> b : A) |  | 
37  | 
(EX f. t=lam x. f(x) \<and> (ALL x. f(x) : A))"  | 
|
| 20140 | 38  | 
apply (unfold HTTgen_def)  | 
39  | 
apply blast  | 
|
40  | 
done  | 
|
41  | 
||
42  | 
lemma HTTXH:  | 
|
| 58977 | 43  | 
"t : HTT \<longleftrightarrow> t=true | t=false | (EX a b. t=<a,b> \<and> a : HTT \<and> b : HTT) |  | 
44  | 
(EX f. t=lam x. f(x) \<and> (ALL x. f(x) : HTT))"  | 
|
| 20140 | 45  | 
apply (rule HTTgen_mono [THEN HTT_def [THEN def_gfp_Tarski], THEN XHlemma1, unfolded HTTgen_def])  | 
46  | 
apply blast  | 
|
47  | 
done  | 
|
48  | 
||
49  | 
||
| 60770 | 50  | 
subsection \<open>Introduction Rules for HTT\<close>  | 
| 20140 | 51  | 
|
| 58977 | 52  | 
lemma HTT_bot: "\<not> bot : HTT"  | 
| 20140 | 53  | 
by (blast dest: HTTXH [THEN iffD1])  | 
54  | 
||
55  | 
lemma HTT_true: "true : HTT"  | 
|
56  | 
by (blast intro: HTTXH [THEN iffD2])  | 
|
57  | 
||
58  | 
lemma HTT_false: "false : HTT"  | 
|
59  | 
by (blast intro: HTTXH [THEN iffD2])  | 
|
60  | 
||
| 58977 | 61  | 
lemma HTT_pair: "<a,b> : HTT \<longleftrightarrow> a : HTT \<and> b : HTT"  | 
| 20140 | 62  | 
apply (rule HTTXH [THEN iff_trans])  | 
63  | 
apply blast  | 
|
64  | 
done  | 
|
65  | 
||
| 58977 | 66  | 
lemma HTT_lam: "lam x. f(x) : HTT \<longleftrightarrow> (ALL x. f(x) : HTT)"  | 
| 20140 | 67  | 
apply (rule HTTXH [THEN iff_trans])  | 
68  | 
apply auto  | 
|
69  | 
done  | 
|
70  | 
||
71  | 
lemmas HTT_rews1 = HTT_bot HTT_true HTT_false HTT_pair HTT_lam  | 
|
72  | 
||
73  | 
lemma HTT_rews2:  | 
|
74  | 
"one : HTT"  | 
|
| 58977 | 75  | 
"inl(a) : HTT \<longleftrightarrow> a : HTT"  | 
76  | 
"inr(b) : HTT \<longleftrightarrow> b : HTT"  | 
|
| 20140 | 77  | 
"zero : HTT"  | 
| 58977 | 78  | 
"succ(n) : HTT \<longleftrightarrow> n : HTT"  | 
| 20140 | 79  | 
"[] : HTT"  | 
| 58977 | 80  | 
"x$xs : HTT \<longleftrightarrow> x : HTT \<and> xs : HTT"  | 
| 20140 | 81  | 
by (simp_all add: data_defs HTT_rews1)  | 
82  | 
||
83  | 
lemmas HTT_rews = HTT_rews1 HTT_rews2  | 
|
84  | 
||
85  | 
||
| 60770 | 86  | 
subsection \<open>Coinduction for HTT\<close>  | 
| 20140 | 87  | 
|
| 58977 | 88  | 
lemma HTT_coinduct: "\<lbrakk>t : R; R <= HTTgen(R)\<rbrakk> \<Longrightarrow> t : HTT"  | 
| 20140 | 89  | 
apply (erule HTT_def [THEN def_coinduct])  | 
90  | 
apply assumption  | 
|
91  | 
done  | 
|
92  | 
||
| 58977 | 93  | 
lemma HTT_coinduct3: "\<lbrakk>t : R; R <= HTTgen(lfp(\<lambda>x. HTTgen(x) Un R Un HTT))\<rbrakk> \<Longrightarrow> t : HTT"  | 
| 20140 | 94  | 
apply (erule HTTgen_mono [THEN [3] HTT_def [THEN def_coinduct3]])  | 
95  | 
apply assumption  | 
|
96  | 
done  | 
|
97  | 
||
| 
32153
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32010 
diff
changeset
 | 
98  | 
lemma HTTgenIs:  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32010 
diff
changeset
 | 
99  | 
"true : HTTgen(R)"  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32010 
diff
changeset
 | 
100  | 
"false : HTTgen(R)"  | 
| 58977 | 101  | 
"\<lbrakk>a : R; b : R\<rbrakk> \<Longrightarrow> <a,b> : HTTgen(R)"  | 
102  | 
"\<And>b. (\<And>x. b(x) : R) \<Longrightarrow> lam x. b(x) : HTTgen(R)"  | 
|
| 
32153
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32010 
diff
changeset
 | 
103  | 
"one : HTTgen(R)"  | 
| 58977 | 104  | 
"a : lfp(\<lambda>x. HTTgen(x) Un R Un HTT) \<Longrightarrow> inl(a) : HTTgen(lfp(\<lambda>x. HTTgen(x) Un R Un HTT))"  | 
105  | 
"b : lfp(\<lambda>x. HTTgen(x) Un R Un HTT) \<Longrightarrow> inr(b) : HTTgen(lfp(\<lambda>x. HTTgen(x) Un R Un HTT))"  | 
|
106  | 
"zero : HTTgen(lfp(\<lambda>x. HTTgen(x) Un R Un HTT))"  | 
|
107  | 
"n : lfp(\<lambda>x. HTTgen(x) Un R Un HTT) \<Longrightarrow> succ(n) : HTTgen(lfp(\<lambda>x. HTTgen(x) Un R Un HTT))"  | 
|
108  | 
"[] : HTTgen(lfp(\<lambda>x. HTTgen(x) Un R Un HTT))"  | 
|
109  | 
"\<lbrakk>h : lfp(\<lambda>x. HTTgen(x) Un R Un HTT); t : lfp(\<lambda>x. HTTgen(x) Un R Un HTT)\<rbrakk> \<Longrightarrow>  | 
|
110  | 
h$t : HTTgen(lfp(\<lambda>x. HTTgen(x) Un R Un HTT))"  | 
|
| 
32153
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32010 
diff
changeset
 | 
111  | 
unfolding data_defs by (genIs HTTgenXH HTTgen_mono)+  | 
| 20140 | 112  | 
|
113  | 
||
| 60770 | 114  | 
subsection \<open>Formation Rules for Types\<close>  | 
| 20140 | 115  | 
|
116  | 
lemma UnitF: "Unit <= HTT"  | 
|
117  | 
by (simp add: subsetXH UnitXH HTT_rews)  | 
|
118  | 
||
119  | 
lemma BoolF: "Bool <= HTT"  | 
|
| 47966 | 120  | 
by (fastforce simp: subsetXH BoolXH iff: HTT_rews)  | 
| 20140 | 121  | 
|
| 58977 | 122  | 
lemma PlusF: "\<lbrakk>A <= HTT; B <= HTT\<rbrakk> \<Longrightarrow> A + B <= HTT"  | 
| 47966 | 123  | 
by (fastforce simp: subsetXH PlusXH iff: HTT_rews)  | 
| 20140 | 124  | 
|
| 58977 | 125  | 
lemma SigmaF: "\<lbrakk>A <= HTT; \<And>x. x:A \<Longrightarrow> B(x) <= HTT\<rbrakk> \<Longrightarrow> SUM x:A. B(x) <= HTT"  | 
| 47966 | 126  | 
by (fastforce simp: subsetXH SgXH HTT_rews)  | 
| 20140 | 127  | 
|
128  | 
||
129  | 
(*** Formation Rules for Recursive types - using coinduction these only need ***)  | 
|
130  | 
(*** exhaution rule for type-former ***)  | 
|
131  | 
||
132  | 
(*Proof by induction - needs induction rule for type*)  | 
|
133  | 
lemma "Nat <= HTT"  | 
|
134  | 
apply (simp add: subsetXH)  | 
|
135  | 
apply clarify  | 
|
136  | 
apply (erule Nat_ind)  | 
|
| 47966 | 137  | 
apply (fastforce iff: HTT_rews)+  | 
| 20140 | 138  | 
done  | 
139  | 
||
140  | 
lemma NatF: "Nat <= HTT"  | 
|
141  | 
apply clarify  | 
|
142  | 
apply (erule HTT_coinduct3)  | 
|
143  | 
apply (fast intro: HTTgenIs elim!: HTTgen_mono [THEN ci3_RI] dest: NatXH [THEN iffD1])  | 
|
144  | 
done  | 
|
145  | 
||
| 58977 | 146  | 
lemma ListF: "A <= HTT \<Longrightarrow> List(A) <= HTT"  | 
| 20140 | 147  | 
apply clarify  | 
148  | 
apply (erule HTT_coinduct3)  | 
|
149  | 
apply (fast intro!: HTTgenIs elim!: HTTgen_mono [THEN ci3_RI]  | 
|
150  | 
subsetD [THEN HTTgen_mono [THEN ci3_AI]]  | 
|
151  | 
dest: ListXH [THEN iffD1])  | 
|
152  | 
done  | 
|
153  | 
||
| 58977 | 154  | 
lemma ListsF: "A <= HTT \<Longrightarrow> Lists(A) <= HTT"  | 
| 20140 | 155  | 
apply clarify  | 
156  | 
apply (erule HTT_coinduct3)  | 
|
157  | 
apply (fast intro!: HTTgenIs elim!: HTTgen_mono [THEN ci3_RI]  | 
|
158  | 
subsetD [THEN HTTgen_mono [THEN ci3_AI]] dest: ListsXH [THEN iffD1])  | 
|
159  | 
done  | 
|
160  | 
||
| 58977 | 161  | 
lemma IListsF: "A <= HTT \<Longrightarrow> ILists(A) <= HTT"  | 
| 20140 | 162  | 
apply clarify  | 
163  | 
apply (erule HTT_coinduct3)  | 
|
164  | 
apply (fast intro!: HTTgenIs elim!: HTTgen_mono [THEN ci3_RI]  | 
|
165  | 
subsetD [THEN HTTgen_mono [THEN ci3_AI]] dest: IListsXH [THEN iffD1])  | 
|
166  | 
done  | 
|
167  | 
||
168  | 
end  |