| author | blanchet | 
| Sun, 01 May 2011 18:37:24 +0200 | |
| changeset 42557 | ae0deb39a254 | 
| parent 40839 | 48e01d16dd17 | 
| child 51247 | 064683ba110c | 
| 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 | 
| 40839 
48e01d16dd17
activate subtyping/coercions in theory Complex_Main;
 wenzelm parents: 
40297diff
changeset | 8 | imports Complex_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 | |
| 40297 | 11 | (* Error messages test *) | 
| 40281 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 12 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 13 | consts func :: "(nat \<Rightarrow> int) \<Rightarrow> nat" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 14 | consts arg :: "int \<Rightarrow> nat" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 15 | (* Invariant arguments | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 16 | term "func arg" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 17 | *) | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 18 | (* No subtype relation - constraint | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 19 | term "(1::nat)::int" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 20 | *) | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 21 | consts func' :: "int \<Rightarrow> int" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 22 | consts arg' :: "nat" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 23 | (* No subtype relation - function application | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 24 | term "func' arg'" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 25 | *) | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 26 | (* Uncomparable types in bound | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 27 | term "arg' = True" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 28 | *) | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 29 | (* Unfullfilled type class requirement | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 30 | term "1 = True" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 31 | *) | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 32 | (* Different constructors | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 33 | term "[1::int] = func" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 34 | *) | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 35 | |
| 40297 | 36 | (* Coercion/type maps definitions *) | 
| 37 | ||
| 40281 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 38 | primrec nat_of_bool :: "bool \<Rightarrow> nat" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 39 | where | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 40 | "nat_of_bool False = 0" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 41 | | "nat_of_bool True = 1" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 42 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 43 | declare [[coercion nat_of_bool]] | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 44 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 45 | declare [[coercion int]] | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 46 | |
| 40297 | 47 | declare [[coercion_map map]] | 
| 40281 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 48 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 49 | definition map_fun :: "('a \<Rightarrow> 'b) \<Rightarrow> ('c \<Rightarrow> 'd) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'd)" where
 | 
| 40282 | 50 | "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 | 51 | |
| 40297 | 52 | 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 | 53 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 54 | primrec map_pair :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a * 'b) \<Rightarrow> ('c * 'd)" where
 | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 55 | "map_pair f g (x,y) = (f x, g y)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 56 | |
| 40297 | 57 | declare [[coercion_map map_pair]] | 
| 40281 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 58 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 59 | (* Examples taken from the haskell draft implementation *) | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 60 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 61 | term "(1::nat) = True" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 62 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 63 | term "True = (1::nat)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 64 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 65 | term "(1::nat) = (True = (1::nat))" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 66 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 67 | term "op = (True = (1::nat))" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 68 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 69 | term "[1::nat,True]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 70 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 71 | term "[True,1::nat]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 72 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 73 | term "[1::nat] = [True]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 74 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 75 | term "[True] = [1::nat]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 76 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 77 | term "[[True]] = [[1::nat]]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 78 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 79 | term "[[[[[[[[[[True]]]]]]]]]] = [[[[[[[[[[1::nat]]]]]]]]]]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 80 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 81 | term "[[True],[42::nat]] = rev [[True]]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 82 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 83 | term "rev [10000::nat] = [False, 420000::nat, True]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 84 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 85 | term "\<lambda> x . x = (3::nat)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 86 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 87 | term "(\<lambda> x . x = (3::nat)) True" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 88 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 89 | term "map (\<lambda> x . x = (3::nat))" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 90 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 91 | 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 | 92 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 93 | consts bnn :: "(bool \<Rightarrow> nat) \<Rightarrow> nat" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 94 | consts nb :: "nat \<Rightarrow> bool" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 95 | consts ab :: "'a \<Rightarrow> bool" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 96 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 97 | term "bnn nb" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 98 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 99 | term "bnn ab" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 100 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 101 | term "\<lambda> x . x = (3::int)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 102 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 103 | term "map (\<lambda> x . x = (3::int)) [True]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 104 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 105 | 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 | 106 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 107 | 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 | 108 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 109 | term "[1::nat,True,1::int,False]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 110 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 111 | 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 | 112 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 113 | consts cbool :: "'a \<Rightarrow> bool" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 114 | consts cnat :: "'a \<Rightarrow> nat" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 115 | consts cint :: "'a \<Rightarrow> int" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 116 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 117 | term "[id, cbool, cnat, cint]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 118 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 119 | 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 | 120 | 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 | 121 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 122 | term "flip funfun" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 123 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 124 | term "map funfun [id,cnat,cint,cbool]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 125 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 126 | term "map (flip funfun True)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 127 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 128 | 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 | 129 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 130 | consts ii :: "int \<Rightarrow> int" | 
| 40282 | 131 | consts aaa :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 40281 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 132 | consts nlist :: "nat list" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 133 | consts ilil :: "int list \<Rightarrow> int list" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 134 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 135 | term "ii (aaa (1::nat) True)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 136 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 137 | term "map ii nlist" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 138 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 139 | term "ilil nlist" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 140 | |
| 
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 | (* Other examples *) | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 144 | |
| 40282 | 145 | definition xs :: "bool list" where "xs = [True]" | 
| 40281 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 146 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 147 | term "(xs::nat list)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 148 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 149 | term "(1::nat) = True" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 150 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 151 | term "True = (1::nat)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 152 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 153 | term "int (1::nat)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 154 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 155 | term "((True::nat)::int)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 156 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 157 | term "1::nat" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 158 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 159 | term "nat 1" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 160 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 161 | definition C :: nat | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 162 | where "C = 123" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 163 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 164 | consts g :: "int \<Rightarrow> int" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 165 | consts h :: "nat \<Rightarrow> nat" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 166 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 167 | term "(g (1::nat)) + (h 2)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 168 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 169 | term "g 1" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 170 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 171 | term "1+(1::nat)" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 172 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 173 | term "((1::int) + (1::nat),(1::int))" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 174 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 175 | 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 | 176 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 177 | term "ys=[[[[[1::nat]]]]]" | 
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 178 | |
| 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 wenzelm parents: diff
changeset | 179 | end |