21 control systems: \<^bold>\<open>Subversion\<close>, \<^bold>\<open>Mercurial\<close>, \<^bold>\<open>Git\<close>. So Phabricator offers |
21 control systems: \<^bold>\<open>Subversion\<close>, \<^bold>\<open>Mercurial\<close>, \<^bold>\<open>Git\<close>. So Phabricator offers |
22 a counter-model to trends of monoculture and centralized version control, |
22 a counter-model to trends of monoculture and centralized version control, |
23 especially due to Microsoft's Github and Atlassian's Bitbucket. |
23 especially due to Microsoft's Github and Atlassian's Bitbucket. |
24 |
24 |
25 The small company behind Phabricator provides paid plans for support and |
25 The small company behind Phabricator provides paid plans for support and |
26 hosting of servers, but it is relatively easy to do \<^emph>\<open>independent |
26 hosting of servers, but it is easy to do \<^emph>\<open>independent self-hosting\<close> on a |
27 self-hosting\<close> on a standard LAMP server (Linux, Apache, MySQL, PHP). This |
27 standard LAMP server (Linux, Apache, MySQL, PHP). This merely requires a |
28 merely requires a virtual Ubuntu server on the net, which can be rented |
28 virtual machine on the Net, which can be rented cheaply from local hosting |
29 rather cheaply from local hosting providers (there is no need to follow big |
29 providers --- there is no need to follow big cloud corporations. So it is |
30 cloud corporations). So it is feasible to remain the master of your virtual |
30 feasible to remain the master of your virtual home, following the slogan |
31 home, following the slogan ``own all your data''. In many respects, |
31 ``own all your data''. In many respects, Phabricator is similar to the |
32 Phabricator is similar to the well-known |
32 well-known Nextcloud\<^footnote>\<open>\<^url>\<open>https://nextcloud.org\<close>\<close> product, concerning both |
33 Nextcloud\<^footnote>\<open>\<^url>\<open>https://nextcloud.com\<close>\<close> product, concerning both the |
33 the technology and sociology. |
34 technology and sociology. |
|
35 |
34 |
36 \<^medskip> |
35 \<^medskip> |
37 The following Phabricator instances may serve as examples: |
36 The following Phabricator instances may serve as examples: |
38 |
37 |
39 \<^item> Phabricator development \<^url>\<open>https://secure.phabricator.com\<close> |
38 \<^item> Phabricator development \<^url>\<open>https://secure.phabricator.com\<close> |
40 \<^item> Wikimedia development \<^url>\<open>https://phabricator.wikimedia.org\<close> |
39 \<^item> Wikimedia development \<^url>\<open>https://phabricator.wikimedia.org\<close> |
41 \<^item> Mercurial development \<^url>\<open>https://phab.mercurial-scm.org\<close> |
40 \<^item> Mercurial development \<^url>\<open>https://phab.mercurial-scm.org\<close> |
42 \<^item> Isabelle development \<^url>\<open>https://isabelle-dev.sketis.net\<close> |
41 \<^item> Isabelle development \<^url>\<open>https://isabelle-dev.sketis.net\<close> |
43 |
42 |
44 \<^medskip> Initial Phabricator server configuration requires many details to be done |
43 \<^medskip> Initial Phabricator configuration requires many details to be done |
45 right.\<^footnote>\<open>See also |
44 right.\<^footnote>\<open>See also |
46 \<^url>\<open>https://secure.phabricator.com/book/phabricator/article/installation_guide\<close> |
45 \<^url>\<open>https://secure.phabricator.com/book/phabricator/article/installation_guide\<close> |
47 in the context of \<^url>\<open>https://help.ubuntu.com/lts/serverguide\<close>.\<close> Isabelle |
46 in the context of \<^url>\<open>https://help.ubuntu.com/lts/serverguide\<close>.\<close> Isabelle |
48 provides some command-line tools to help with the setup, and afterwards |
47 provides some command-line tools to help with the setup, and afterwards |
49 Isabelle support is optional: it is possible to run and maintain the server, |
48 Isabelle support is optional: it is possible to run and maintain the server, |
50 without requiring a full Isabelle distribution again. |
49 without requiring the somewhat bulky Isabelle distribution again. |
51 \<close> |
50 \<close> |
52 |
51 |
53 |
52 |
54 section \<open>Quick start\<close> |
53 section \<open>Quick start\<close> |
55 |
54 |
56 text \<open> |
55 text \<open> |
57 The starting point is a fresh installation of \<^bold>\<open>Ubuntu 18.04 LTS\<close>: that |
56 The starting point is a fresh installation of \<^bold>\<open>Ubuntu 18.04 |
58 particular version is required due to subtle dependencies on system |
57 LTS\<close>\<^footnote>\<open>\<^url>\<open>https://ubuntu.com/download\<close>\<close>: this version is mandatory due to |
59 configuration and add-on packages. |
58 subtle dependencies on system packages and configuration that is assumed by |
|
59 the Isabelle setup tool. |
60 |
60 |
61 For production use, a proper \<^emph>\<open>Virtual Server\<close> or \<^emph>\<open>Root Server\<close> product |
61 For production use, a proper \<^emph>\<open>Virtual Server\<close> or \<^emph>\<open>Root Server\<close> product |
62 from a hosting provider will be required, including an Internet Domain Name |
62 from a hosting provider will be required, including an Internet Domain Name |
63 (a pseudo sub-domain in the style of Apache is sufficient). |
63 (a pseudo sub-domain in the style of Apache is sufficient). |
64 |
64 |
65 Initial experimentation also works on a local host, e.g.\ via |
65 Initial experimentation also works on a local host, e.g.\ via |
66 VirtualBox\<^footnote>\<open>\<^url>\<open>https://www.virtualbox.org\<close>\<close>. The public server behind |
66 VirtualBox\<^footnote>\<open>\<^url>\<open>https://www.virtualbox.org\<close>\<close>. The public domain \<^verbatim>\<open>lvh.me\<close> is |
67 \<^verbatim>\<open>lvh.me\<close> provides arbitrary subdomains as an alias to \<^verbatim>\<open>localhost\<close>, which |
67 used below: it maps arbitrary subdomains to \<^verbatim>\<open>localhost\<close>. |
68 will be used for the default installation below. |
|
69 |
68 |
70 All administrative commands need to be run as \<^verbatim>\<open>root\<close> user (e.g.\ via |
69 All administrative commands need to be run as \<^verbatim>\<open>root\<close> user (e.g.\ via |
71 \<^verbatim>\<open>sudo\<close>). |
70 \<^verbatim>\<open>sudo\<close>). Note that Isabelle refers to user-specific configuration in the |
|
71 user home directory via @{setting ISABELLE_HOME_USER} |
|
72 (\secref{sec:settings}); that may be different or absent for the root user |
|
73 and thus cause confusion. |
72 \<close> |
74 \<close> |
73 |
75 |
74 |
76 |
75 subsection \<open>Initial setup\<close> |
77 subsection \<open>Initial setup\<close> |
76 |
78 |
77 text \<open> |
79 text \<open> |
78 Isabelle can managed multiple named installations Phabricator installations: |
80 Isabelle can manage multiple named Phabricator installations: this allows to |
79 this allows to separate administrative responsibilities, e.g.\ different |
81 separate administrative responsibilities, e.g.\ different approaches to user |
80 approaches to user management for different projects. Subsequently we always |
82 management for different projects. Subsequently we always use the default |
81 use the implicit default name ``\<^verbatim>\<open>vcs\<close>'': it will appear in file and |
83 name ``\<^verbatim>\<open>vcs\<close>'': it will appear in file and directory locations, internal |
82 directory locations, internal database names and URLs. |
84 database names and URLs. |
83 |
85 |
84 The initial setup (with full Linux package upgrade) works as follows: |
86 The initial setup works as follows (with full Linux package upgrade): |
85 |
87 |
86 @{verbatim [display] \<open> isabelle phabricator_setup -U\<close>} |
88 @{verbatim [display] \<open> isabelle phabricator_setup -U\<close>} |
87 |
89 |
88 After installing many packages and cloning the Phabricator distribution, the |
90 After installing many packages, cloning the Phabricator distribution, |
89 final output of the tool should be the URL for further manual configuration |
91 initializing the MySQL database and Apache, the tool prints an URL for |
90 (using a local web browser). Here the key points are: |
92 further configuration. The following needs to be provided by the web |
91 |
93 interface: |
92 \<^item> An initial user account that will get administrator rights. There is no |
94 |
93 need to create a separate \<^verbatim>\<open>admin\<close> account. Instead, a regular user that |
95 \<^item> An initial user that will get administrator rights. There is no need to |
94 will take over this responsibility can be used here. Subsequently we |
96 create a special \<^verbatim>\<open>admin\<close> account. Instead, a regular user that will take |
95 assume that user \<^verbatim>\<open>makarius\<close> becomes the initial administrator. |
97 over this responsibility can be used here. Subsequently we assume that |
|
98 user \<^verbatim>\<open>makarius\<close> becomes the initial administrator. |
96 |
99 |
97 \<^item> An \<^emph>\<open>Auth Provider\<close> to manage user names and passwords. None is provided |
100 \<^item> An \<^emph>\<open>Auth Provider\<close> to manage user names and passwords. None is provided |
98 by default, and Phabricator points out this omission prominently in its |
101 by default, and Phabricator points out this omission prominently in its |
99 overview of \<^emph>\<open>Setup Issues\<close>: following these hints quickly leads to the |
102 overview of \<^emph>\<open>Setup Issues\<close>: following these hints quickly leads to the |
100 place where a regular \<^emph>\<open>Username/Password\<close> provider can be added. |
103 place where a regular \<^emph>\<open>Username/Password\<close> provider can be added. |
101 |
104 |
102 Note that Phabricator allows to delegate the responsibility of |
105 Alternatively, Phabricator can delegate the responsibility of |
103 authentication to big corporations like Google and Facebook, but these can |
106 authentication to big corporations like Google and Facebook, but these can |
104 be easily avoided. Genuine self-hosting means to manage users directly, |
107 be ignored. Genuine self-hosting means to manage users directly, without |
105 without outsourcing of authentication. |
108 outsourcing of authentication. |
106 |
109 |
107 \<^medskip> |
110 \<^item> A proper password for the administrator can now be set, e.g.\ by the |
108 With the Auth Provider available, the administrator can now set a proper |
111 following command: |
109 password. This works e.g.\ by initiating a local password reset procedure: |
112 |
110 |
113 @{verbatim [display] \<open> isabelle phabricator bin/auth recover makarius\<close>} |
111 @{verbatim [display] \<open> isabelle phabricator bin/auth recover makarius\<close>} |
114 |
112 |
115 The printed URL gives access to a login and password dialog in the web |
113 The printed URL gives access to a password dialog in the web browser. |
116 interface. |
114 |
117 |
115 Further users will be able to provide a password more directly, because the |
118 Any further users will be able to provide a password directly, because the |
116 Auth Provider is now active. |
119 Auth Provider is already active. |
117 |
120 |
118 \<^medskip> |
121 \<^item> The list of Phabricator \<^bold>\<open>Setup Issues\<close> should be studied with some |
119 The pending request in Phabricator \<^bold>\<open>Setup Issues\<close> to lock the configuration |
122 care, to make sure that no serious problems are remaining. The request to |
120 can be fulfilled as follows: |
123 lock the configuration can be fulfilled as follows: |
121 |
124 |
122 @{verbatim [display] \<open> isabelle phabricator bin/auth lock\<close>} |
125 @{verbatim [display] \<open> isabelle phabricator bin/auth lock\<close>} |
123 |
126 |
124 \<^medskip> |
127 \<^medskip> A few other Setup Issues might be relevant as well, e.g.\ the timezone |
125 Most other Setup Issues can be ignored, after reading through them briefly |
128 of the server. Some more exotic points can be ignored: Phabricator |
126 to make sure that there are no genuine problems remaining. |
129 provides careful explanations about what it thinks could be wrong, always |
|
130 leaving some room for interpretation. |
127 \<close> |
131 \<close> |
128 |
132 |
129 |
133 |
130 subsection \<open>Mailer configuration\<close> |
134 subsection \<open>Mailer configuration\<close> |
131 |
135 |
132 text \<open> |
136 text \<open> |
133 The next important thing is messaging: Phabricator needs to be able to |
137 The next important thing is messaging: Phabricator needs to be able to |
134 communicate with users. There are many variations on \<^emph>\<open>Mailer |
138 communicate with users on its own account, e.g.\ to reset passwords. The |
135 Configuration\<close>, but a conventional SMTP server connection with a dedicated |
139 documentation has many variations on \<^emph>\<open>Configuring Outbound |
136 \<^verbatim>\<open>phabricator\<close> user is sufficient. Note that there is no need to run a mail |
140 Email\<close>\<^footnote>\<open>\<^url>\<open>https://secure.phabricator.com/book/phabricator/article/configuring_outbound_email\<close>\<close>, |
137 server on the self-hosted Linux machine: regular web-hosting packages |
141 but a conventional SMTP server with a dedicated \<^verbatim>\<open>phabricator\<close> user is |
138 usually allow to create new mail accounts easily. (As a last resort it is |
142 sufficient. There is no need to run a separate mail server on the |
139 possible to use a service like Gmail, but that would again introduce |
143 self-hosted Linux machine: hosting providers often include such a service |
140 unnecessary dependency on Google.) |
144 for free, e.g.\ as part of a web-hosting package. As a last resort it is |
|
145 also possible to use a corporate service like Gmail, but such dependency |
|
146 dilutes the whole effort of self-hosting. |
141 |
147 |
142 \<^medskip> |
148 \<^medskip> |
143 Mailer configuration requires a few command-line invocations as follows: |
149 Mailer configuration requires a few command-line invocations as follows: |
144 |
150 |
145 @{verbatim [display] \<open> isabelle phabricator_setup_mail\<close>} |
151 @{verbatim [display] \<open> isabelle phabricator_setup_mail\<close>} |
146 |
152 |
147 \<^noindent> Now edit the generated JSON file to provide the mail account details. Then |
153 \<^noindent> This generates a JSON template file for the the mail account details. |
148 add and test it with Phabricator like this (to let Phabricator send a |
154 After editing that, the subsequent command will add and test it with |
149 message to the administrator from above): |
155 Phabricator: |
150 |
156 |
151 @{verbatim [display] \<open> isabelle phabricator_setup_mail -T makarius\<close>} |
157 @{verbatim [display] \<open> isabelle phabricator_setup_mail -T makarius\<close>} |
152 |
158 |
153 The mail configuration process can be refined and repeated until it really |
159 This tells Phabricator to send a message to the administrator created |
154 works: host name, port number, protocol etc.\ all need to be correct. The |
160 before; the output informs about success or errors. |
155 \<^verbatim>\<open>key\<close> field in the JSON file identifies the name of the configuration; it |
161 |
|
162 The mail configuration process can be refined and repeated until it works |
|
163 properly: host name, port number, protocol etc.\ all need to be correct. The |
|
164 \<^verbatim>\<open>key\<close> field in the JSON file identifies the name of the configuration that |
156 will be overwritten each time, when taking over the parameters via |
165 will be overwritten each time, when taking over the parameters via |
157 \<^verbatim>\<open>isabelle phabricator_setup_mail\<close>. |
166 \<^verbatim>\<open>isabelle phabricator_setup_mail\<close>. |
158 |
167 |
159 \<^medskip> |
168 \<^medskip> |
160 For information, the resulting mailer configuration can be queried like |
169 This is how to query the current mail configuration: |
161 this: |
|
162 |
170 |
163 @{verbatim [display] \<open> isabelle phabricator bin/config get cluster.mailers\<close>} |
171 @{verbatim [display] \<open> isabelle phabricator bin/config get cluster.mailers\<close>} |
164 \<close> |
172 \<close> |
165 |
173 |
166 |
174 |
167 section \<open>Reference of command-line tools\<close> |
175 section \<open>Reference of command-line tools\<close> |
168 |
176 |
169 text \<open> |
177 text \<open> |
170 The subsequent command-line tools usually require root user privileges on |
178 The subsequent command-line tools usually require root user privileges on |
171 the underlying Linux system (e.g.\ via \<^verbatim>\<open>sudo bash\<close> to open a subshell, or |
179 the underlying Linux system (e.g.\ via \<^verbatim>\<open>sudo bash\<close> to open a subshell, or |
172 directly via \<^verbatim>\<open>sudo isabelle ...\<close>). Note that Isabelle refers to |
180 directly via \<^verbatim>\<open>sudo isabelle phabricator ...\<close>). |
173 user-specific configuration in the user home directory via @{setting |
|
174 ISABELLE_HOME_USER} (\secref{sec:settings}); that may be different or absent |
|
175 for the root user and thus cause confusion. |
|
176 \<close> |
181 \<close> |
177 |
182 |
178 |
183 |
179 subsection \<open>\<^verbatim>\<open>isabelle phabricator\<close>\<close> |
184 subsection \<open>\<^verbatim>\<open>isabelle phabricator\<close>\<close> |
180 |
185 |
238 Install Phabricator as LAMP application (Linux, Apache, MySQL, PHP). |
243 Install Phabricator as LAMP application (Linux, Apache, MySQL, PHP). |
239 |
244 |
240 The installation name (default: "vcs") is mapped to a regular |
245 The installation name (default: "vcs") is mapped to a regular |
241 Unix user; this is relevant for public SSH access.\<close>} |
246 Unix user; this is relevant for public SSH access.\<close>} |
242 |
247 |
243 Installation requires Linux root user access. All required packages are |
248 Installation requires Linux root permissions. All required packages are |
244 installed automatically beforehand, this includes the Apache web server and |
249 installed automatically beforehand, this includes the Apache web server and |
245 the MySQL database engine. |
250 the MySQL database engine. |
246 |
251 |
247 Global configuration in \<^verbatim>\<open>/etc\<close> and other directories like \<^verbatim>\<open>/var/www\<close> |
252 Global configuration in \<^verbatim>\<open>/etc\<close> or a few other directories like \<^verbatim>\<open>/var/www\<close> |
248 always use name prefixes \<^verbatim>\<open>isabelle-phabricator\<close> or \<^verbatim>\<open>phabricator\<close>. Local |
253 uses name prefixes like \<^verbatim>\<open>isabelle-phabricator\<close> or \<^verbatim>\<open>phabricator\<close>. Local |
249 configuration for a particular installation uses more specific names derived |
254 configuration for a particular installation uses more specific names derived |
250 from \<^verbatim>\<open>phabricator-\<close>\<open>NAME\<close>, e.g. \<^verbatim>\<open>/var/www/phabricator-vcs\<close> for a default |
255 from \<^verbatim>\<open>phabricator-\<close>\<open>NAME\<close>, e.g.\ \<^verbatim>\<open>/var/www/phabricator-vcs\<close> for the |
251 installation. |
256 default. |
252 |
257 |
253 Knowing the conventions, it is possible to purge a Linux installation from |
258 Knowing the naming conventions, it is possible to purge a Linux installation |
254 Isabelle/Phabricator with some effort. There is no automated |
259 from Isabelle/Phabricator with some effort, but there is no automated |
255 de-installation: it is often better to re-install a clean virtual machine |
260 procedure for de-installation. In the worst case, it might be better to |
256 image. |
261 re-install the virtual machine from a clean image. |
257 |
262 |
258 \<^medskip> |
263 \<^medskip> |
259 Option \<^verbatim>\<open>-U\<close> ensures a full update of system packages, before installing |
264 Option \<^verbatim>\<open>-U\<close> ensures a full update of system packages, before installing |
260 further packages required by Phabricator. |
265 further packages required by Phabricator. This might require to reboot. |
261 |
266 |
262 Option \<^verbatim>\<open>-n\<close> provides an alternative installation name. The default name |
267 Option \<^verbatim>\<open>-n\<close> provides an alternative installation name. The default name |
263 \<^verbatim>\<open>vcs\<close> means ``Version Control System''. The name will appear in the URL for |
268 \<^verbatim>\<open>vcs\<close> means ``version control system''. The name appears in the URL for SSH |
264 SSH access, and thus has some relevance to end-users. The initial server URL |
269 access, and thus has some relevance to end-users. The initial server URL |
265 also uses that name suffix, but it can be changed later via regular Apache |
270 also uses the same suffix, but that can (and should) be changed later via |
266 configuration. |
271 regular Apache configuration. |
267 |
272 |
268 Option \<^verbatim>\<open>-r\<close> specifies an alternative installation root directory: it needs |
273 Option \<^verbatim>\<open>-r\<close> specifies an alternative installation root directory: it needs |
269 to be accessible for the Apache web server. |
274 to be accessible for the Apache web server. |
270 |
275 |
271 Option \<^verbatim>\<open>-R\<close> specifies an alternative directory for repositories that are |
276 Option \<^verbatim>\<open>-R\<close> specifies an alternative directory for repositories that are |
272 hosted by Phabricator. Provided that it is accessible for the Apache web |
277 hosted by Phabricator. Provided that it is accessible for the Apache web |
273 server, the directory can be re-used directly for the builtin \<^verbatim>\<open>hgweb\<close> view |
278 server, the directory can be reused for the \<^verbatim>\<open>hgweb\<close> view by Mercurial.\<^footnote>\<open>See |
274 by Mercurial. See also the documentation |
279 also the documentation |
275 \<^url>\<open>https://www.mercurial-scm.org/wiki/PublishingRepositories#hgweb\<close> and the |
280 \<^url>\<open>https://www.mercurial-scm.org/wiki/PublishingRepositories\<close> and the |
276 example \<^url>\<open>https://isabelle.sketis.net/repos\<close> for repositories in |
281 example \<^url>\<open>https://isabelle.sketis.net/repos\<close>.\<close> |
277 \<^url>\<open>https://isabelle-dev.sketis.net/diffusion\<close>. |
|
278 \<close> |
282 \<close> |
279 |
283 |
280 |
284 |
281 subsection \<open>\<^verbatim>\<open>isabelle phabricator_setup_mail\<close>\<close> |
285 subsection \<open>\<^verbatim>\<open>isabelle phabricator_setup_mail\<close>\<close> |
282 |
286 |
283 text \<open> |
287 text \<open> |
284 The @{tool_def phabricator_setup_mail} provides mail configuration for an |
288 The @{tool_def phabricator_setup_mail} provides mail configuration for an |
285 existing Phabricator installation (preferably via SMTP): |
289 existing Phabricator installation: |
286 @{verbatim [display] \<open>Usage: isabelle phabricator_setup_mail [OPTIONS] |
290 @{verbatim [display] \<open>Usage: isabelle phabricator_setup_mail [OPTIONS] |
287 |
291 |
288 Options are: |
292 Options are: |
289 -T USER send test mail to Phabricator user |
293 -T USER send test mail to Phabricator user |
290 -f FILE config file (default: "mailers.json" within |
294 -f FILE config file (default: "mailers.json" within |