src/HOLCF/Tutorial/New_Domain.thy
author huffman
Sat, 30 Oct 2010 15:13:11 -0700
changeset 40329 73f2b99b549d
parent 39986 38677db30cad
child 40497 d2e876d6da8c
permissions -rw-r--r--
change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     1
(*  Title:      HOLCF/ex/New_Domain.thy
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     3
*)
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     4
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     5
header {* Definitional domain package *}
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     6
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     7
theory New_Domain
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     8
imports HOLCF
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
     9
begin
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    10
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    11
text {*
40329
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 39986
diff changeset
    12
  UPDATE: The definitional back-end is now the default mode of the domain
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 39986
diff changeset
    13
  package. This file should be merged with @{text Domain_ex.thy}.
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    14
*}
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    15
39986
38677db30cad rename class 'sfp' to 'bifinite'
huffman
parents: 39974
diff changeset
    16
default_sort bifinite
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    17
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    18
text {*
39986
38677db30cad rename class 'sfp' to 'bifinite'
huffman
parents: 39974
diff changeset
    19
  Provided that @{text bifinite} is the default sort, the @{text new_domain}
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    20
  package should work with any type definition supported by the old
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    21
  domain package.
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    22
*}
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    23
40329
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 39986
diff changeset
    24
domain 'a llist = LNil | LCons (lazy 'a) (lazy "'a llist")
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    25
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    26
text {*
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    27
  The difference is that the new domain package is completely
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    28
  definitional, and does not generate any axioms.  The following type
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    29
  and constant definitions are not produced by the old domain package.
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    30
*}
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    31
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    32
thm type_definition_llist
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    33
thm llist_abs_def llist_rep_def
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    34
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    35
text {*
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    36
  The new domain package also adds support for indirect recursion with
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    37
  user-defined datatypes.  This definition of a tree datatype uses
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    38
  indirect recursion through the lazy list type constructor.
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    39
*}
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    40
40329
73f2b99b549d change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
huffman
parents: 39986
diff changeset
    41
domain 'a ltree = Leaf (lazy 'a) | Branch (lazy "'a ltree llist")
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    42
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    43
text {*
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    44
  For indirect-recursive definitions, the domain package is not able to
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    45
  generate a high-level induction rule.  (It produces a warning
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    46
  message instead.)  The low-level reach lemma (now proved as a
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    47
  theorem, no longer generated as an axiom) can be used to derive
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    48
  other induction rules.
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    49
*}
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    50
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    51
thm ltree.reach
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    52
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    53
text {*
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    54
  The definition of the take function uses map functions associated with
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    55
  each type constructor involved in the definition.  A map function
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    56
  for the lazy list type has been generated by the new domain package.
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    57
*}
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    58
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    59
thm ltree.take_rews
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    60
thm llist_map_def
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    61
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    62
lemma ltree_induct:
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    63
  fixes P :: "'a ltree \<Rightarrow> bool"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    64
  assumes adm: "adm P"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    65
  assumes bot: "P \<bottom>"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    66
  assumes Leaf: "\<And>x. P (Leaf\<cdot>x)"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    67
  assumes Branch: "\<And>f l. \<forall>x. P (f\<cdot>x) \<Longrightarrow> P (Branch\<cdot>(llist_map\<cdot>f\<cdot>l))"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    68
  shows "P x"
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    69
proof -
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    70
  have "P (\<Squnion>i. ltree_take i\<cdot>x)"
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    71
  using adm
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    72
  proof (rule admD)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    73
    fix i
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    74
    show "P (ltree_take i\<cdot>x)"
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    75
    proof (induct i arbitrary: x)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    76
      case (0 x)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    77
      show "P (ltree_take 0\<cdot>x)" by (simp add: bot)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    78
    next
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    79
      case (Suc n x)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    80
      show "P (ltree_take (Suc n)\<cdot>x)"
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    81
        apply (cases x)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    82
        apply (simp add: bot)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    83
        apply (simp add: Leaf)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    84
        apply (simp add: Branch Suc)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    85
        done
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    86
    qed
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 33813
diff changeset
    87
  qed (simp add: ltree.chain_take)
33813
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    88
  thus ?thesis
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    89
    by (simp add: ltree.reach)
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    90
qed
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    91
0bc8d4f786bd example theory for new domain package
huffman
parents:
diff changeset
    92
end