| author | blanchet | 
| Thu, 11 Sep 2014 18:54:36 +0200 | |
| changeset 58299 | 30ab8289f0e1 | 
| parent 55932 | 68c5104d2204 | 
| child 67399 | eab6ce8368fa | 
| permissions | -rw-r--r-- | 
| 40282 | 1  | 
(* Title: HOL/ex/Coercion_Examples.thy  | 
2  | 
Author: Dmitriy Traytel, TU Muenchen  | 
|
3  | 
||
4  | 
Examples for coercive subtyping via subtype constraints.  | 
|
5  | 
*)  | 
|
6  | 
||
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
theory Coercion_Examples  | 
| 
51247
 
064683ba110c
Coercion_Examples defines required coercions itself (no Complex_Main needed)
 
traytel 
parents: 
40839 
diff
changeset
 | 
8  | 
imports Main  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
begin  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
|
| 
51247
 
064683ba110c
Coercion_Examples defines required coercions itself (no Complex_Main needed)
 
traytel 
parents: 
40839 
diff
changeset
 | 
11  | 
declare[[coercion_enabled]]  | 
| 
 
064683ba110c
Coercion_Examples defines required coercions itself (no Complex_Main needed)
 
traytel 
parents: 
40839 
diff
changeset
 | 
12  | 
|
| 40297 | 13  | 
(* Error messages test *)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
15  | 
consts func :: "(nat \<Rightarrow> int) \<Rightarrow> nat"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
consts arg :: "int \<Rightarrow> nat"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
(* Invariant arguments  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
term "func arg"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
(* No subtype relation - constraint  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
term "(1::nat)::int"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
consts func' :: "int \<Rightarrow> int"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
consts arg' :: "nat"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
(* No subtype relation - function application  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
term "func' arg'"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
(* Uncomparable types in bound  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
term "arg' = True"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
(* Unfullfilled type class requirement  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
32  | 
term "1 = True"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
(* Different constructors  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
term "[1::int] = func"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
|
| 40297 | 38  | 
(* Coercion/type maps definitions *)  | 
39  | 
||
| 54438 | 40  | 
abbreviation nat_of_bool :: "bool \<Rightarrow> nat"  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
41  | 
where  | 
| 54438 | 42  | 
"nat_of_bool \<equiv> of_bool"  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
44  | 
declare [[coercion nat_of_bool]]  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
declare [[coercion int]]  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
|
| 40297 | 48  | 
declare [[coercion_map map]]  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
definition map_fun :: "('a \<Rightarrow> 'b) \<Rightarrow> ('c \<Rightarrow> 'd) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'd)" where
 | 
| 40282 | 51  | 
"map_fun f g h = g o h o f"  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
|
| 40297 | 53  | 
declare [[coercion_map "\<lambda> f g h . g o h o f"]]  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
54  | 
|
| 55932 | 55  | 
primrec map_prod :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a * 'b) \<Rightarrow> ('c * 'd)" where
 | 
56  | 
"map_prod f g (x,y) = (f x, g y)"  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
|
| 55932 | 58  | 
declare [[coercion_map map_prod]]  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
60  | 
(* Examples taken from the haskell draft implementation *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
61  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
62  | 
term "(1::nat) = True"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
63  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
64  | 
term "True = (1::nat)"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
65  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
66  | 
term "(1::nat) = (True = (1::nat))"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
67  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
68  | 
term "op = (True = (1::nat))"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
70  | 
term "[1::nat,True]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
71  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
72  | 
term "[True,1::nat]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
74  | 
term "[1::nat] = [True]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
75  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
76  | 
term "[True] = [1::nat]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
78  | 
term "[[True]] = [[1::nat]]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
79  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
80  | 
term "[[[[[[[[[[True]]]]]]]]]] = [[[[[[[[[[1::nat]]]]]]]]]]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
81  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
82  | 
term "[[True],[42::nat]] = rev [[True]]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
83  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
84  | 
term "rev [10000::nat] = [False, 420000::nat, True]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
85  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
86  | 
term "\<lambda> x . x = (3::nat)"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
87  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
88  | 
term "(\<lambda> x . x = (3::nat)) True"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
90  | 
term "map (\<lambda> x . x = (3::nat))"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
91  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
92  | 
term "map (\<lambda> x . x = (3::nat)) [True,1::nat]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
94  | 
consts bnn :: "(bool \<Rightarrow> nat) \<Rightarrow> nat"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
95  | 
consts nb :: "nat \<Rightarrow> bool"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
96  | 
consts ab :: "'a \<Rightarrow> bool"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
97  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
term "bnn nb"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
100  | 
term "bnn ab"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
101  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
term "\<lambda> x . x = (3::int)"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
term "map (\<lambda> x . x = (3::int)) [True]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
105  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
term "map (\<lambda> x . x = (3::int)) [True,1::nat]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
107  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
term "map (\<lambda> x . x = (3::int)) [True,1::nat,1::int]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
110  | 
term "[1::nat,True,1::int,False]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
111  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
112  | 
term "map (map (\<lambda> x . x = (3::int))) [[True],[1::nat],[True,1::int]]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
113  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
114  | 
consts cbool :: "'a \<Rightarrow> bool"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
115  | 
consts cnat :: "'a \<Rightarrow> nat"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
116  | 
consts cint :: "'a \<Rightarrow> int"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
117  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
118  | 
term "[id, cbool, cnat, cint]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
119  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
120  | 
consts funfun :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
121  | 
consts flip :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c"
 | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
122  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
123  | 
term "flip funfun"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
124  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
125  | 
term "map funfun [id,cnat,cint,cbool]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
126  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
127  | 
term "map (flip funfun True)"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
128  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
129  | 
term "map (flip funfun True) [id,cnat,cint,cbool]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
130  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
131  | 
consts ii :: "int \<Rightarrow> int"  | 
| 40282 | 132  | 
consts aaa :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
133  | 
consts nlist :: "nat list"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
134  | 
consts ilil :: "int list \<Rightarrow> int list"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
135  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
term "ii (aaa (1::nat) True)"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
137  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
138  | 
term "map ii nlist"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
139  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
140  | 
term "ilil nlist"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
141  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
142  | 
(***************************************************)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
143  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
144  | 
(* Other examples *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
145  | 
|
| 40282 | 146  | 
definition xs :: "bool list" where "xs = [True]"  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
147  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
148  | 
term "(xs::nat list)"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
150  | 
term "(1::nat) = True"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
151  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
152  | 
term "True = (1::nat)"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
153  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
154  | 
term "int (1::nat)"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
155  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
156  | 
term "((True::nat)::int)"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
157  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
158  | 
term "1::nat"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
159  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
160  | 
term "nat 1"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
161  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
162  | 
definition C :: nat  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
163  | 
where "C = 123"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
164  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
165  | 
consts g :: "int \<Rightarrow> int"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
166  | 
consts h :: "nat \<Rightarrow> nat"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
167  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
168  | 
term "(g (1::nat)) + (h 2)"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
169  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
170  | 
term "g 1"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
171  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
172  | 
term "1+(1::nat)"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
173  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
174  | 
term "((1::int) + (1::nat),(1::int))"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
175  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
176  | 
definition ys :: "bool list list list list list" where "ys=[[[[[True]]]]]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
177  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
178  | 
term "ys=[[[[[1::nat]]]]]"  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
179  | 
|
| 51335 | 180  | 
typedecl ('a, 'b, 'c) F
 | 
181  | 
consts Fmap :: "('a \<Rightarrow> 'd) \<Rightarrow> ('a, 'b, 'c) F \<Rightarrow> ('d, 'b, 'c) F"
 | 
|
182  | 
consts z :: "(bool, nat, bool) F"  | 
|
183  | 
declare [[coercion_map "Fmap :: ('a \<Rightarrow> 'd) \<Rightarrow> ('a, 'b, 'c) F \<Rightarrow> ('d, 'b, 'c) F"]]
 | 
|
184  | 
term "z :: (nat, nat, bool) F"  | 
|
185  | 
||
| 51320 | 186  | 
consts  | 
187  | 
case_nil :: "'a \<Rightarrow> 'b"  | 
|
188  | 
  case_cons :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
|
189  | 
  case_abs :: "('c \<Rightarrow> 'b) \<Rightarrow> 'b"
 | 
|
190  | 
case_elem :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b"  | 
|
191  | 
||
192  | 
declare [[coercion_args case_cons - -]]  | 
|
193  | 
declare [[coercion_args case_abs -]]  | 
|
194  | 
declare [[coercion_args case_elem - +]]  | 
|
195  | 
||
196  | 
term "case_cons (case_abs (\<lambda>n. case_abs (\<lambda>is. case_elem (((n::nat),(is::int list))) (n#is)))) case_nil"  | 
|
197  | 
||
198  | 
consts n :: nat m :: nat  | 
|
199  | 
term "- (n + m)"  | 
|
200  | 
declare [[coercion_args uminus -]]  | 
|
201  | 
declare [[coercion_args plus + +]]  | 
|
202  | 
term "- (n + m)"  | 
|
| 54438 | 203  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
204  | 
end  |