src/HOLCF/ex/Domain_ex.thy
author huffman
Mon, 12 Apr 2010 19:29:16 -0700
changeset 36120 dd6e69cdcc1e
parent 35781 b7738ab762b1
permissions -rw-r--r--
update domain package examples
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
     1
(*  Title:      HOLCF/ex/Domain_ex.thy
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
     3
*)
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
     4
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
     5
header {* Domain package examples *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
     6
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
     7
theory Domain_ex
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
     8
imports HOLCF
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
     9
begin
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    10
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    11
text {* Domain constructors are strict by default. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    12
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    13
domain d1 = d1a | d1b "d1" "d1"
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    14
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    15
lemma "d1b\<cdot>\<bottom>\<cdot>y = \<bottom>" by simp
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    16
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    17
text {* Constructors can be made lazy using the @{text "lazy"} keyword. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    18
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    19
domain d2 = d2a | d2b (lazy "d2")
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    20
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    21
lemma "d2b\<cdot>x \<noteq> \<bottom>" by simp
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    22
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    23
text {* Strict and lazy arguments may be mixed arbitrarily. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    24
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    25
domain d3 = d3a | d3b (lazy "d2") "d2"
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    26
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    27
lemma "P (d3b\<cdot>x\<cdot>y = \<bottom>) \<longleftrightarrow> P (y = \<bottom>)" by simp
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    28
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    29
text {* Selectors can be used with strict or lazy constructor arguments. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    30
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    31
domain d4 = d4a | d4b (lazy d4b_left :: "d2") (d4b_right :: "d2")
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    32
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    33
lemma "y \<noteq> \<bottom> \<Longrightarrow> d4b_left\<cdot>(d4b\<cdot>x\<cdot>y) = x" by simp
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    34
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    35
text {* Mixfix declarations can be given for data constructors. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    36
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    37
domain d5 = d5a | d5b (lazy "d5") "d5" (infixl ":#:" 70)
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    38
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    39
lemma "d5a \<noteq> x :#: y :#: z" by simp
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    40
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    41
text {* Mixfix declarations can also be given for type constructors. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    42
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    43
domain ('a, 'b) lazypair (infixl ":*:" 25) =
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    44
  lpair (lazy lfst :: 'a) (lazy lsnd :: 'b) (infixl ":*:" 75)
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    45
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    46
lemma "\<forall>p::('a :*: 'b). p \<sqsubseteq> lfst\<cdot>p :*: lsnd\<cdot>p"
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    47
by (rule allI, case_tac p, simp_all)
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    48
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    49
text {* Non-recursive constructor arguments can have arbitrary types. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    50
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    51
domain ('a, 'b) d6 = d6 "int lift" "'a \<oplus> 'b u" (lazy "('a :*: 'b) \<times> ('b \<rightarrow> 'a)")
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    52
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    53
text {*
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    54
  Indirect recusion is allowed for sums, products, lifting, and the
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35497
diff changeset
    55
  continuous function space.  However, the domain package does not
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35497
diff changeset
    56
  generate an induction rule in terms of the constructors.
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    57
*}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    58
31232
689aa7da48cc define copy functions using combinators; add checking for failed proofs of induction rules
huffman
parents: 30922
diff changeset
    59
domain 'a d7 = d7a "'a d7 \<oplus> int lift" | d7b "'a \<otimes> 'a d7" | d7c (lazy "'a d7 \<rightarrow> 'a")
36120
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
    60
  -- "Indirect recursion detected, skipping proofs of (co)induction rules"
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
    61
(* d7.induct is absent *)
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    62
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    63
text {*
36120
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
    64
  Indirect recursion is also allowed using previously-defined datatypes.
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    65
*}
36120
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
    66
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    67
domain 'a slist = SNil | SCons 'a "'a slist"
36120
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
    68
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
    69
domain 'a stree = STip | SBranch "'a stree slist"
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    70
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    71
text {* Mutually-recursive datatypes can be defined using the @{text "and"} keyword. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    72
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    73
domain d8 = d8a | d8b "d9" and d9 = d9a | d9b (lazy "d8")
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    74
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    75
text {* Non-regular recursion is not allowed. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    76
(*
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    77
domain ('a, 'b) altlist = ANil | ACons 'a "('b, 'a) altlist"
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    78
  -- "illegal direct recursion with different arguments"
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    79
domain 'a nest = Nest1 'a | Nest2 "'a nest nest"
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    80
  -- "illegal direct recursion with different arguments"
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    81
*)
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    82
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    83
text {*
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    84
  Mutually-recursive datatypes must have all the same type arguments,
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    85
  not necessarily in the same order.
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    86
*}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    87
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    88
domain ('a, 'b) list1 = Nil1 | Cons1 'a "('b, 'a) list2"
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    89
   and ('b, 'a) list2 = Nil2 | Cons2 'b "('a, 'b) list1"
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    90
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    91
text {* Induction rules for flat datatypes have no admissibility side-condition. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    92
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    93
domain 'a flattree = Tip | Branch "'a flattree" "'a flattree"
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    94
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    95
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"
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35661
diff changeset
    96
by (rule flattree.induct) -- "no admissibility requirement"
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    97
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    98
text {* Trivial datatypes will produce a warning message. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
    99
35443
2e0f9516947e change domain package's treatment of variable names in theorems to be like datatype package
huffman
parents: 31232
diff changeset
   100
domain triv = Triv triv triv
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   101
  -- "domain Domain_ex.triv is empty!"
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   102
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   103
lemma "(x::triv) = \<bottom>" by (induct x, simp_all)
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   104
36120
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
   105
text {* Lazy constructor arguments may have unpointed types. *}
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
   106
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
   107
domain natlist = nnil | ncons (lazy "nat discr") natlist
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
   108
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
   109
text {* Class constraints may be given for type parameters on the LHS. *}
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
   110
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
   111
domain ('a::cpo) box = Box (lazy 'a)
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
   112
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   113
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   114
subsection {* Generated constants and theorems *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   115
35661
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35642
diff changeset
   116
domain 'a tree = Leaf (lazy 'a) | Node (left :: "'a tree") (right :: "'a tree")
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   117
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   118
lemmas tree_abs_defined_iff =
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   119
  iso.abs_defined_iff [OF iso.intro [OF tree.abs_iso tree.rep_iso]]
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   120
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   121
text {* Rules about ismorphism *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   122
term tree_rep
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   123
term tree_abs
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   124
thm tree.rep_iso
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   125
thm tree.abs_iso
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   126
thm tree.iso_rews
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   127
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   128
text {* Rules about constructors *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   129
term Leaf
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   130
term Node
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents: 35443
diff changeset
   131
thm Leaf_def Node_def
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35661
diff changeset
   132
thm tree.nchotomy
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   133
thm tree.exhaust
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   134
thm tree.compacts
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   135
thm tree.con_rews
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   136
thm tree.dist_les
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   137
thm tree.dist_eqs
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   138
thm tree.inverts
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   139
thm tree.injects
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   140
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   141
text {* Rules about case combinator *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   142
term tree_when
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35444
diff changeset
   143
thm tree.tree_when_def
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   144
thm tree.when_rews
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   145
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   146
text {* Rules about selectors *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   147
term left
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   148
term right
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   149
thm tree.sel_rews
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   150
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   151
text {* Rules about discriminators *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   152
term is_Leaf
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   153
term is_Node
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   154
thm tree.dis_rews
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   155
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   156
text {* Rules about pattern match combinators *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   157
term Leaf_pat
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   158
term Node_pat
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   159
thm tree.pat_rews
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   160
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   161
text {* Rules about monadic pattern match combinators *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   162
term match_Leaf
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   163
term match_Node
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   164
thm tree.match_rews
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   165
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   166
text {* Rules about take function *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   167
term tree_take
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   168
thm tree.take_def
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35444
diff changeset
   169
thm tree.take_0
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35444
diff changeset
   170
thm tree.take_Suc
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   171
thm tree.take_rews
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35444
diff changeset
   172
thm tree.chain_take
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35444
diff changeset
   173
thm tree.take_take
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35444
diff changeset
   174
thm tree.deflation_take
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35661
diff changeset
   175
thm tree.take_below
35642
f478d5a9d238 generate separate qualified theorem name for each type's reach and take_lemma
huffman
parents: 35585
diff changeset
   176
thm tree.take_lemma
35585
555f26f00e47 skip proof of induction rule for indirect-recursive domain definitions
huffman
parents: 35497
diff changeset
   177
thm tree.lub_take
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35444
diff changeset
   178
thm tree.reach
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35661
diff changeset
   179
thm tree.finite_induct
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   180
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   181
text {* Rules about finiteness predicate *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   182
term tree_finite
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   183
thm tree.finite_def
35661
ede27eb8e94b don't generate rule foo.finites for non-flat domains; use take_induct rule to prove induction rule
huffman
parents: 35642
diff changeset
   184
thm tree.finite (* only generated for flat datatypes *)
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   185
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   186
text {* Rules about bisimulation predicate *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   187
term tree_bisim
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   188
thm tree.bisim_def
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35661
diff changeset
   189
thm tree.coinduct
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   190
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   191
text {* Induction rule *}
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35661
diff changeset
   192
thm tree.induct
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   193
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   194
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   195
subsection {* Known bugs *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   196
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   197
text {* Declaring a mixfix with spaces causes some strange parse errors. *}
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   198
(*
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   199
domain xx = xx ("x y")
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   200
  -- "Inner syntax error: unexpected end of input"
30922
96d053e00ec0 add more examples to Domain_ex.thy
huffman
parents: 30920
diff changeset
   201
*)
96d053e00ec0 add more examples to Domain_ex.thy
huffman
parents: 30920
diff changeset
   202
96d053e00ec0 add more examples to Domain_ex.thy
huffman
parents: 30920
diff changeset
   203
text {*
36120
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
   204
  Non-cpo type parameters currently do not work.
30922
96d053e00ec0 add more examples to Domain_ex.thy
huffman
parents: 30920
diff changeset
   205
*}
96d053e00ec0 add more examples to Domain_ex.thy
huffman
parents: 30920
diff changeset
   206
(*
36120
dd6e69cdcc1e update domain package examples
huffman
parents: 35781
diff changeset
   207
domain ('a::type) stream = snil | scons (lazy "'a discr") "'a stream"
30922
96d053e00ec0 add more examples to Domain_ex.thy
huffman
parents: 30920
diff changeset
   208
*)
96d053e00ec0 add more examples to Domain_ex.thy
huffman
parents: 30920
diff changeset
   209
30920
811ab0923a62 add HOLCF/ex/Domain_ex.thy, with example uses of the domain package
huffman
parents:
diff changeset
   210
end