src/HOLCF/ex/Domain_ex.thy
changeset 30920 811ab0923a62
child 30922 96d053e00ec0
equal deleted inserted replaced
30919:dcf8a7a66bd1 30920:811ab0923a62
       
     1 (*  Title:      HOLCF/ex/Domain_ex.thy
       
     2     Author:     Brian Huffman
       
     3 *)
       
     4 
       
     5 header {* Domain package examples *}
       
     6 
       
     7 theory Domain_ex
       
     8 imports HOLCF
       
     9 begin
       
    10 
       
    11 text {* Domain constructors are strict by default. *}
       
    12 
       
    13 domain d1 = d1a | d1b "d1" "d1"
       
    14 
       
    15 lemma "d1b\<cdot>\<bottom>\<cdot>y = \<bottom>" by simp
       
    16 
       
    17 text {* Constructors can be made lazy using the @{text "lazy"} keyword. *}
       
    18 
       
    19 domain d2 = d2a | d2b (lazy "d2")
       
    20 
       
    21 lemma "d2b\<cdot>x \<noteq> \<bottom>" by simp
       
    22 
       
    23 text {* Strict and lazy arguments may be mixed arbitrarily. *}
       
    24 
       
    25 domain d3 = d3a | d3b (lazy "d2") "d2"
       
    26 
       
    27 lemma "P (d3b\<cdot>x\<cdot>y = \<bottom>) \<longleftrightarrow> P (y = \<bottom>)" by simp
       
    28 
       
    29 text {* Selectors can be used with strict or lazy constructor arguments. *}
       
    30 
       
    31 domain d4 = d4a | d4b (lazy d4b_left :: "d2") (d4b_right :: "d2")
       
    32 
       
    33 lemma "y \<noteq> \<bottom> \<Longrightarrow> d4b_left\<cdot>(d4b\<cdot>x\<cdot>y) = x" by simp
       
    34 
       
    35 text {* Mixfix declarations can be given for data constructors. *}
       
    36 
       
    37 domain d5 = d5a | d5b (lazy "d5") "d5" (infixl ":#:" 70)
       
    38 
       
    39 lemma "d5a \<noteq> x :#: y :#: z" by simp
       
    40 
       
    41 text {* Mixfix declarations can also be given for type constructors. *}
       
    42 
       
    43 domain ('a, 'b) lazypair (infixl ":*:" 25) =
       
    44   lpair (lazy lfst :: 'a) (lazy lsnd :: 'b) (infixl ":*:" 75)
       
    45 
       
    46 lemma "\<forall>p::('a :*: 'b). p \<sqsubseteq> lfst\<cdot>p :*: lsnd\<cdot>p"
       
    47 by (rule allI, case_tac p, simp_all)
       
    48 
       
    49 text {* Non-recursive constructor arguments can have arbitrary types. *}
       
    50 
       
    51 domain ('a, 'b) d6 = d6 "int lift" "'a \<oplus> 'b u" (lazy "('a :*: 'b) \<times> ('b \<rightarrow> 'a)")
       
    52 
       
    53 text {*
       
    54   Indirect recusion is allowed for sums, products, lifting, and the
       
    55   continuous function space.  However, the domain package currently
       
    56   generates induction rules that are too weak.  A fix is planned for
       
    57   the next release.
       
    58 *}
       
    59 
       
    60 domain 'a d7 = d7a "'a d7 \<oplus> int lift" | d7b "'a \<otimes> 'a d7" | d7c "'a d7 \<rightarrow> 'a"
       
    61 
       
    62 thm d7.ind -- "note the lack of inductive hypotheses"
       
    63 
       
    64 text {*
       
    65   Indirect recursion using previously-defined datatypes is currently
       
    66   not allowed.  This restriction should go away by the next release.
       
    67 *}
       
    68 (*
       
    69 domain 'a slist = SNil | SCons 'a "'a slist"
       
    70 domain 'a stree = STip | SBranch "'a stree slist" -- "illegal indirect recursion"
       
    71 *)
       
    72 
       
    73 text {* Mutually-recursive datatypes can be defined using the @{text "and"} keyword. *}
       
    74 
       
    75 domain d8 = d8a | d8b "d9" and d9 = d9a | d9b (lazy "d8")
       
    76 
       
    77 text {* Non-regular recursion is not allowed. *}
       
    78 (*
       
    79 domain ('a, 'b) altlist = ANil | ACons 'a "('b, 'a) altlist"
       
    80   -- "illegal direct recursion with different arguments"
       
    81 domain 'a nest = Nest1 'a | Nest2 "'a nest nest"
       
    82   -- "illegal direct recursion with different arguments"
       
    83 *)
       
    84 
       
    85 text {*
       
    86   Mutually-recursive datatypes must have all the same type arguments,
       
    87   not necessarily in the same order.
       
    88 *}
       
    89 
       
    90 domain ('a, 'b) list1 = Nil1 | Cons1 'a "('b, 'a) list2"
       
    91    and ('b, 'a) list2 = Nil2 | Cons2 'b "('a, 'b) list1"
       
    92 
       
    93 text {* Induction rules for flat datatypes have no admissibility side-condition. *}
       
    94 
       
    95 domain 'a flattree = Tip | Branch "'a flattree" "'a flattree"
       
    96 
       
    97 lemma "\<lbrakk>P \<bottom>; P Tip; \<And>x y. \<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>; P x; P y\<rbrakk> \<Longrightarrow> P (Branch\<cdot>x\<cdot>y)\<rbrakk> \<Longrightarrow> P x"
       
    98 by (rule flattree.ind) -- "no admissibility requirement"
       
    99 
       
   100 text {* Trivial datatypes will produce a warning message. *}
       
   101 
       
   102 domain triv = triv1 triv triv
       
   103   -- "domain Domain_ex.triv is empty!"
       
   104 
       
   105 lemma "(x::triv) = \<bottom>" by (induct x, simp_all)
       
   106 
       
   107 
       
   108 subsection {* Generated constants and theorems *}
       
   109 
       
   110 domain 'a tree = Leaf (lazy 'a) | Node (left :: "'a tree") (lazy right :: "'a tree")
       
   111 
       
   112 lemmas tree_abs_defined_iff =
       
   113   iso.abs_defined_iff [OF iso.intro [OF tree.abs_iso tree.rep_iso]]
       
   114 
       
   115 text {* Rules about ismorphism *}
       
   116 term tree_rep
       
   117 term tree_abs
       
   118 thm tree.rep_iso
       
   119 thm tree.abs_iso
       
   120 thm tree.iso_rews
       
   121 
       
   122 text {* Rules about constructors *}
       
   123 term Leaf
       
   124 term Node
       
   125 thm tree.Leaf_def tree.Node_def
       
   126 thm tree.exhaust
       
   127 thm tree.casedist
       
   128 thm tree.compacts
       
   129 thm tree.con_rews
       
   130 thm tree.dist_les
       
   131 thm tree.dist_eqs
       
   132 thm tree.inverts
       
   133 thm tree.injects
       
   134 
       
   135 text {* Rules about case combinator *}
       
   136 term tree_when
       
   137 thm tree.when_def
       
   138 thm tree.when_rews
       
   139 
       
   140 text {* Rules about selectors *}
       
   141 term left
       
   142 term right
       
   143 thm tree.sel_rews
       
   144 
       
   145 text {* Rules about discriminators *}
       
   146 term is_Leaf
       
   147 term is_Node
       
   148 thm tree.dis_rews
       
   149 
       
   150 text {* Rules about pattern match combinators *}
       
   151 term Leaf_pat
       
   152 term Node_pat
       
   153 thm tree.pat_rews
       
   154 
       
   155 text {* Rules about monadic pattern match combinators *}
       
   156 term match_Leaf
       
   157 term match_Node
       
   158 thm tree.match_rews
       
   159 
       
   160 text {* Rules about copy function *}
       
   161 term tree_copy
       
   162 thm tree.copy_def
       
   163 thm tree.copy_rews
       
   164 
       
   165 text {* Rules about take function *}
       
   166 term tree_take
       
   167 thm tree.take_def
       
   168 thm tree.take_rews
       
   169 thm tree.take_lemmas
       
   170 thm tree.finite_ind
       
   171 
       
   172 text {* Rules about finiteness predicate *}
       
   173 term tree_finite
       
   174 thm tree.finite_def
       
   175 thm tree.finites
       
   176 
       
   177 text {* Rules about bisimulation predicate *}
       
   178 term tree_bisim
       
   179 thm tree.bisim_def
       
   180 thm tree.coind
       
   181 
       
   182 text {* Induction rule *}
       
   183 thm tree.ind
       
   184 
       
   185 
       
   186 subsection {* Known bugs *}
       
   187 
       
   188 text {* Declaring a mixfix with spaces causes some strange parse errors. *}
       
   189 (*
       
   190 domain xx = xx ("x y")
       
   191   -- "Inner syntax error: unexpected end of input"
       
   192 
       
   193 domain 'a foo = foo (sel::"'a") ("a b")
       
   194   -- {* Inner syntax error at "= UU" *}
       
   195 *)
       
   196 
       
   197 text {*
       
   198   I don't know what is going on here.  The failed proof has to do with
       
   199   the finiteness predicate.
       
   200 *}
       
   201 (*
       
   202 domain foo = Foo (lazy "bar") and bar = Bar
       
   203   -- "Tactic failed."
       
   204 *)
       
   205 
       
   206 end